From 10721bab946e62d158fd3b8fa88667e1f851180c Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Fri, 6 May 2022 19:31:54 +0000 Subject: [PATCH 01/37] feat(topology/algebra/module/basic): basic topological properties of quotient modules (#13433) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit More precisely, we prove that : * if `M` is a topological module and `S` is a submodule of `M`, then `M ⧸ S` is a topological module * furthermore, if `S` is closed, then `M ⧸ S` is regular (hence T2) - [x] depends on: #13278 - [x] depends on: #13401 --- src/topology/algebra/module/basic.lean | 39 ++++++++++++++++++++++++++ 1 file changed, 39 insertions(+) diff --git a/src/topology/algebra/module/basic.lean b/src/topology/algebra/module/basic.lean index 8b654b46a1707..361cc657498cd 100644 --- a/src/topology/algebra/module/basic.lean +++ b/src/topology/algebra/module/basic.lean @@ -1917,3 +1917,42 @@ lemma continuous_linear_map.closed_complemented_ker_of_right_inverse {R : Type*} (h : function.right_inverse f₂ f₁) : f₁.ker.closed_complemented := ⟨f₁.proj_ker_of_right_inverse f₂ h, f₁.proj_ker_of_right_inverse_apply_idem f₂ h⟩ + +section quotient + +namespace submodule + +variables {R M : Type*} [ring R] [add_comm_group M] [module R M] [topological_space M] + (S : submodule R M) + +lemma is_open_map_mkq [topological_add_group M] : is_open_map S.mkq := +quotient_add_group.is_open_map_coe S.to_add_subgroup + +instance topological_add_group_quotient [topological_add_group M] : + topological_add_group (M ⧸ S) := +topological_add_group_quotient S.to_add_subgroup + +instance has_continuous_smul_quotient [topological_space R] [topological_add_group M] + [has_continuous_smul R M] : + has_continuous_smul R (M ⧸ S) := +begin + split, + have quot : quotient_map (λ au : R × M, (au.1, S.mkq au.2)), + from is_open_map.to_quotient_map + (is_open_map.id.prod S.is_open_map_mkq) + (continuous_id.prod_map continuous_quot_mk) + (function.surjective_id.prod_map $ surjective_quot_mk _), + rw quot.continuous_iff, + exact continuous_quot_mk.comp continuous_smul +end + +instance regular_quotient_of_is_closed [topological_add_group M] [is_closed (S : set M)] : + regular_space (M ⧸ S) := +begin + letI : is_closed (S.to_add_subgroup : set M) := ‹_›, + exact S.to_add_subgroup.regular_quotient_of_is_closed +end + +end submodule + +end quotient From dcd452deeda31bb585d4c7527396a0ffb8fc2f7c Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Fri, 6 May 2022 20:39:04 +0000 Subject: [PATCH 02/37] =?UTF-8?q?feat(analysis/locally=5Fconvex/with=5Fsem?= =?UTF-8?q?inorms):=20characterization=20of=20the=20topology=20induced=20b?= =?UTF-8?q?y=20seminorms=20in=20terms=20of=20`=F0=9D=93=9D=200`=20(#13547)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This shows that a topology is induced by the family of seminorms `p` iff `𝓝 0 = ⨅ i, (𝓝 0).comap (p i)`, which allows to use the extensive filter and topology library (see e.g. #13549). --- .../locally_convex/with_seminorms.lean | 27 +++++++++++++++++++ src/analysis/seminorm.lean | 7 +++++ 2 files changed, 34 insertions(+) diff --git a/src/analysis/locally_convex/with_seminorms.lean b/src/analysis/locally_convex/with_seminorms.lean index 08e0e2e7ae3d2..12bef187ce7d9 100644 --- a/src/analysis/locally_convex/with_seminorms.lean +++ b/src/analysis/locally_convex/with_seminorms.lean @@ -172,6 +172,25 @@ protected def module_filter_basis : module_filter_basis 𝕜 E := smul_left' := p.basis_sets_smul_left, smul_right' := p.basis_sets_smul_right } +lemma filter_eq_infi (p : seminorm_family 𝕜 E ι) : + p.module_filter_basis.to_filter_basis.filter = ⨅ i, (𝓝 0).comap (p i) := +begin + refine le_antisymm (le_infi $ λ i, _) _, + { rw p.module_filter_basis.to_filter_basis.has_basis.le_basis_iff + (metric.nhds_basis_ball.comap _), + intros ε hε, + refine ⟨(p i).ball 0 ε, _, _⟩, + { rw ← (finset.sup_singleton : _ = p i), + exact p.basis_sets_mem {i} hε, }, + { rw [id, (p i).ball_zero_eq_preimage_ball] } }, + { rw p.module_filter_basis.to_filter_basis.has_basis.ge_iff, + rintros U (hU : U ∈ p.basis_sets), + rcases p.basis_sets_iff.mp hU with ⟨s, r, hr, rfl⟩, + rw [id, seminorm.ball_finset_sup_eq_Inter _ _ _ hr, s.Inter_mem_sets], + exact λ i hi, filter.mem_infi_of_mem i ⟨metric.ball 0 r, metric.ball_mem_nhds 0 hr, + eq.subset ((p i).ball_zero_eq_preimage_ball).symm⟩, }, +end + end seminorm_family end filter_basis @@ -282,6 +301,14 @@ lemma seminorm_family.with_seminorms_of_has_basis (p : seminorm_family 𝕜 E ι p.with_seminorms_of_nhds $ filter.has_basis.eq_of_same_basis h p.add_group_filter_basis.to_filter_basis.has_basis +lemma seminorm_family.with_seminorms_iff_nhds_eq_infi (p : seminorm_family 𝕜 E ι) : + with_seminorms p ↔ (𝓝 0 : filter E) = ⨅ i, (𝓝 0).comap (p i) := +begin + rw ← p.filter_eq_infi, + refine ⟨λ h, _, p.with_seminorms_of_nhds⟩, + rw h.topology_eq_with_seminorms, + exact add_group_filter_basis.nhds_zero_eq _, +end end topological_add_group diff --git a/src/analysis/seminorm.lean b/src/analysis/seminorm.lean index 08f07a5dcb3a5..6d3d784dff027 100644 --- a/src/analysis/seminorm.lean +++ b/src/analysis/seminorm.lean @@ -470,6 +470,13 @@ end section norm_one_class variables [norm_one_class 𝕜] (p : seminorm 𝕜 E) +lemma ball_zero_eq_preimage_ball {r : ℝ} : + p.ball 0 r = p ⁻¹' (metric.ball 0 r) := +begin + ext x, + simp only [mem_ball, sub_zero, mem_preimage, mem_ball_zero_iff, real.norm_of_nonneg (p.nonneg x)], +end + @[simp] lemma ball_bot {r : ℝ} (x : E) (hr : 0 < r) : ball (⊥ : seminorm 𝕜 E) x r = set.univ := ball_zero' x hr From 275dabfc96302f0d4d973cd206e16ef2910e6bce Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sat, 7 May 2022 00:43:46 +0000 Subject: [PATCH 03/37] feat(order/atoms): is_atomic_of_order_bot_lt_well_founded (#13967) Co-authored-by: Scott Morrison --- src/order/atoms.lean | 14 ++++++++++++++ 1 file changed, 14 insertions(+) diff --git a/src/order/atoms.lean b/src/order/atoms.lean index bce44154fb480..67113c02b8548 100644 --- a/src/order/atoms.lean +++ b/src/order/atoms.lean @@ -196,6 +196,20 @@ theorem is_coatomic_iff_forall_is_coatomic_Ici [order_top α] : is_atomic_dual_iff_is_coatomic.symm.trans $ is_atomic_iff_forall_is_atomic_Iic.trans $ forall_congr (λ x, is_coatomic_dual_iff_is_atomic.symm.trans iff.rfl) +section well_founded + +lemma is_atomic_of_order_bot_well_founded_lt [order_bot α] + (h : well_founded ((<) : α → α → Prop)) : is_atomic α := +⟨λ a, or_iff_not_imp_left.2 $ + λ ha, let ⟨b, hb, hm⟩ := h.has_min { b | b ≠ ⊥ ∧ b ≤ a } ⟨a, ha, le_rfl⟩ in + ⟨b, ⟨hb.1, λ c, not_imp_not.1 $ λ hc hl, hm c ⟨hc, hl.le.trans hb.2⟩ hl⟩, hb.2⟩⟩ + +lemma is_coatomic_of_order_top_gt_well_founded [order_top α] + (h : well_founded ((>) : α → α → Prop)) : is_coatomic α := +is_atomic_dual_iff_is_coatomic.1 (@is_atomic_of_order_bot_well_founded_lt αᵒᵈ _ _ h) + +end well_founded + end atomic section atomistic From 5789c637d5855d9fe2094bda16be177d4f36477e Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Sat, 7 May 2022 04:10:25 +0000 Subject: [PATCH 04/37] chore(scripts): update nolints.txt (#14004) 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 3551a7d157a1f..90505df64e22a 100644 --- a/scripts/nolints.txt +++ b/scripts/nolints.txt @@ -493,13 +493,9 @@ apply_nolint relator.right_total doc_blame apply_nolint relator.right_unique doc_blame -- measure_theory/group/prod.lean -apply_nolint measure_theory.absolutely_continuous_of_is_mul_left_invariant to_additive_doc 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_eq to_additive_doc -apply_nolint measure_theory.map_prod_mul_eq_swap to_additive_doc apply_nolint measure_theory.map_prod_mul_inv_eq to_additive_doc -apply_nolint measure_theory.measure_eq_div_smul 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 From ca1375a5a53a12877cf366316abff5cac77fc4af Mon Sep 17 00:00:00 2001 From: negiizhao Date: Sat, 7 May 2022 04:45:56 +0000 Subject: [PATCH 05/37] refactor(algebra/order/monoid_lemmas): reorder the file (#13492) Just like in `algebra/order/monoid_lemmas_zero_lt`, sort by algebraic assumptions and order assumptions first, then put similar lemmas together. It would be simpler to find duplicates, missing lemmas, and inconsistencies. (There are so many!) --- src/algebra/order/monoid_lemmas.lean | 751 +++++++++++++-------------- 1 file changed, 369 insertions(+), 382 deletions(-) diff --git a/src/algebra/order/monoid_lemmas.lean b/src/algebra/order/monoid_lemmas.lean index 91c1474073fb0..4808acd20a01f 100644 --- a/src/algebra/order/monoid_lemmas.lean +++ b/src/algebra/order/monoid_lemmas.lean @@ -1,7 +1,8 @@ /- 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, Damiano Testa +Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl, Damiano Testa, +Yuyang Zhao -/ import algebra.covariant_and_contravariant import order.monotone @@ -40,7 +41,8 @@ variables [has_le α] /- The prime on this lemma is present only on the multiplicative version. The unprimed version is taken by the analogous lemma for semiring, with an extra non-negativity assumption. -/ @[to_additive add_le_add_left] -lemma mul_le_mul_left' [covariant_class α α (*) (≤)] {b c : α} (bc : b ≤ c) (a : α) : +lemma mul_le_mul_left' [covariant_class α α (*) (≤)] + {b c : α} (bc : b ≤ c) (a : α) : a * b ≤ a * c := covariant_class.elim _ bc @@ -96,7 +98,8 @@ lemma mul_lt_mul_iff_right rel_iff_cov α α (swap (*)) (<) a @[to_additive add_lt_add_left] -lemma mul_lt_mul_left' [covariant_class α α (*) (<)] {b c : α} (bc : b < c) (a : α) : +lemma mul_lt_mul_left' [covariant_class α α (*) (<)] + {b c : α} (bc : b < c) (a : α) : a * b < a * c := covariant_class.elim _ bc @@ -120,6 +123,108 @@ contravariant_class.elim a bc end has_lt +section preorder +variables [preorder α] + +@[to_additive] +lemma mul_lt_mul_of_lt_of_lt [covariant_class α α (*) (<)] [covariant_class α α (swap (*)) (<)] + {a b c d : α} (h₁ : a < b) (h₂ : c < d) : a * c < b * d := +calc a * c < a * d : mul_lt_mul_left' h₂ a + ... < b * d : mul_lt_mul_right' h₁ d + +@[to_additive] +lemma mul_lt_mul_of_le_of_lt [covariant_class α α (*) (<)] [covariant_class α α (swap (*)) (≤)] + {a b c d : α} (h₁ : a ≤ b) (h₂ : c < d) : a * c < b * d := +(mul_le_mul_right' h₁ _).trans_lt (mul_lt_mul_left' h₂ b) + +@[to_additive] +lemma mul_lt_mul_of_lt_of_le [covariant_class α α (*) (≤)] [covariant_class α α (swap (*)) (<)] + {a b c d : α} (h₁ : a < b) (h₂ : c ≤ d) : a * c < b * d := +(mul_le_mul_left' h₂ _).trans_lt (mul_lt_mul_right' h₁ d) + +@[to_additive add_lt_add] +lemma mul_lt_mul''' [covariant_class α α (*) (<)] [covariant_class α α (swap (*)) (≤)] + {a b c d : α} (h₁ : a < b) (h₂ : c < d) : a * c < b * d := +mul_lt_mul_of_le_of_lt h₁.le h₂ + +@[to_additive add_le_add] +lemma mul_le_mul' [covariant_class α α (*) (≤)] [covariant_class α α (swap (*)) (≤)] + {a b c d : α} (h₁ : a ≤ b) (h₂ : c ≤ d) : a * c ≤ b * d := +(mul_le_mul_left' h₂ _).trans (mul_le_mul_right' h₁ d) + +@[to_additive] +lemma mul_le_mul_three [covariant_class α α (*) (≤)] [covariant_class α α (swap (*)) (≤)] + {a b c d e f : α} (h₁ : a ≤ d) (h₂ : b ≤ e) (h₃ : c ≤ f) : + a * b * c ≤ d * e * f := +mul_le_mul' (mul_le_mul' h₁ h₂) h₃ + +@[to_additive] +lemma mul_lt_of_mul_lt_left [covariant_class α α (*) (≤)] + {a b c d : α} (h : a * b < c) (hle : d ≤ b) : + a * d < c := +(mul_le_mul_left' hle a).trans_lt h + +@[to_additive] +lemma mul_le_of_mul_le_left [covariant_class α α (*) (≤)] + {a b c d : α} (h : a * b ≤ c) (hle : d ≤ b) : + a * d ≤ c := +@act_rel_of_rel_of_act_rel _ _ _ (≤) _ ⟨λ _ _ _, le_trans⟩ a _ _ _ hle h + +@[to_additive] +lemma mul_lt_of_mul_lt_right [covariant_class α α (swap (*)) (≤)] + {a b c d : α} (h : a * b < c) (hle : d ≤ a) : + d * b < c := +(mul_le_mul_right' hle b).trans_lt h + +@[to_additive] +lemma mul_le_of_mul_le_right [covariant_class α α (swap (*)) (≤)] + {a b c d : α} (h : a * b ≤ c) (hle : d ≤ a) : + d * b ≤ c := +(mul_le_mul_right' hle b).trans h + +@[to_additive] +lemma lt_mul_of_lt_mul_left [covariant_class α α (*) (≤)] + {a b c d : α} (h : a < b * c) (hle : c ≤ d) : + a < b * d := +h.trans_le (mul_le_mul_left' hle b) + +@[to_additive] +lemma le_mul_of_le_mul_left [covariant_class α α (*) (≤)] + {a b c d : α} (h : a ≤ b * c) (hle : c ≤ d) : + a ≤ b * d := +@rel_act_of_rel_of_rel_act _ _ _ (≤) _ ⟨λ _ _ _, le_trans⟩ b _ _ _ hle h + +@[to_additive] +lemma lt_mul_of_lt_mul_right [covariant_class α α (swap (*)) (≤)] + {a b c d : α} (h : a < b * c) (hle : b ≤ d) : + a < d * c := +h.trans_le (mul_le_mul_right' hle c) + +@[to_additive] +lemma le_mul_of_le_mul_right [covariant_class α α (swap (*)) (≤)] + {a b c d : α} (h : a ≤ b * c) (hle : b ≤ d) : + a ≤ d * c := +h.trans (mul_le_mul_right' hle c) + +end preorder + +section partial_order +variables [partial_order α] + +@[to_additive] +lemma mul_left_cancel'' [contravariant_class α α (*) (≤)] + {a b c : α} (h : a * b = a * c) : + b = c := +(le_of_mul_le_mul_left' h.le).antisymm (le_of_mul_le_mul_left' h.ge) + +@[to_additive] +lemma mul_right_cancel'' [contravariant_class α α (swap (*)) (≤)] + {a b c : α} (h : a * b = c * b) : + a = c := +le_antisymm (le_of_mul_le_mul_right' h.le) (le_of_mul_le_mul_right' h.ge) + +end partial_order + end has_mul -- using one @@ -129,17 +234,23 @@ variables [mul_one_class α] section has_le variables [has_le α] +@[to_additive le_add_of_nonneg_right] +lemma le_mul_of_one_le_right' [covariant_class α α (*) (≤)] {a b : α} (h : 1 ≤ b) : + a ≤ a * b := +calc a = a * 1 : (mul_one _).symm + ... ≤ a * b : mul_le_mul_left' h a + +@[to_additive add_le_of_nonpos_right] +lemma mul_le_of_le_one_right' [covariant_class α α (*) (≤)] {a b : α} (h : b ≤ 1) : + a * b ≤ a := +calc a * b ≤ a * 1 : mul_le_mul_left' h a + ... = a : mul_one a + @[simp, to_additive le_add_iff_nonneg_right] lemma le_mul_iff_one_le_right' - [covariant_class α α (*) (≤)] [contravariant_class α α (*) (≤)] - (a : α) {b : α} : a ≤ a * b ↔ 1 ≤ b := -iff.trans (by rw [mul_one]) (mul_le_mul_iff_left a) - -@[simp, to_additive add_le_iff_nonpos_right] -lemma mul_le_iff_le_one_right' [covariant_class α α (*) (≤)] [contravariant_class α α (*) (≤)] (a : α) {b : α} : - a * b ≤ a ↔ b ≤ 1 := + a ≤ a * b ↔ 1 ≤ b := iff.trans (by rw [mul_one]) (mul_le_mul_iff_left a) @[simp, to_additive le_add_iff_nonneg_left] @@ -149,6 +260,13 @@ lemma le_mul_iff_one_le_left' a ≤ b * a ↔ 1 ≤ b := iff.trans (by rw one_mul) (mul_le_mul_iff_right a) +@[simp, to_additive add_le_iff_nonpos_right] +lemma mul_le_iff_le_one_right' + [covariant_class α α (*) (≤)] [contravariant_class α α (*) (≤)] + (a : α) {b : α} : + a * b ≤ a ↔ b ≤ 1 := +iff.trans (by rw [mul_one]) (mul_le_mul_iff_left a) + @[simp, to_additive add_le_iff_nonpos_left] lemma mul_le_iff_le_one_left' [covariant_class α α (swap (*)) (≤)] [contravariant_class α α (swap (*)) (≤)] @@ -158,27 +276,14 @@ iff.trans (by rw one_mul) (mul_le_mul_iff_right b) end has_le -lemma exists_square_le {α : Type*} [mul_one_class α] [linear_order α] [covariant_class α α (*) (<)] - (a : α) : ∃ (b : α), b * b ≤ a := -begin - by_cases h : a < 1, - { use a, - have : a*a < a*1, - exact mul_lt_mul_left' h a, - rw mul_one at this, - exact le_of_lt this }, - { use 1, - push_neg at h, - rwa mul_one } -end - section has_lt variable [has_lt α] @[to_additive lt_add_of_pos_right] lemma lt_mul_of_one_lt_right' [covariant_class α α (*) (<)] - (a : α) {b : α} (h : 1 < b) : a < a * b := + (a : α) {b : α} (h : 1 < b) : + a < a * b := calc a = a * 1 : (mul_one _).symm ... < a * b : mul_lt_mul_left' h a @@ -189,18 +294,20 @@ lemma lt_mul_iff_one_lt_right' a < a * b ↔ 1 < b := iff.trans (by rw mul_one) (mul_lt_mul_iff_left a) -@[simp, to_additive add_lt_iff_neg_left] -lemma mul_lt_iff_lt_one_left' - [covariant_class α α (*) (<)] [contravariant_class α α (*) (<)] {a b : α} : - a * b < a ↔ b < 1 := -iff.trans (by rw mul_one) (mul_lt_mul_iff_left a) - @[simp, to_additive lt_add_iff_pos_left] lemma lt_mul_iff_one_lt_left' [covariant_class α α (swap (*)) (<)] [contravariant_class α α (swap (*)) (<)] - (a : α) {b : α} : a < b * a ↔ 1 < b := + (a : α) {b : α} : + a < b * a ↔ 1 < b := iff.trans (by rw one_mul) (mul_lt_mul_iff_right a) +@[simp, to_additive add_lt_iff_neg_left] +lemma mul_lt_iff_lt_one_left' + [covariant_class α α (*) (<)] [contravariant_class α α (*) (<)] + {a b : α} : + a * b < a ↔ b < 1 := +iff.trans (by rw mul_one) (mul_lt_mul_iff_left a) + @[simp, to_additive add_lt_iff_neg_right] lemma mul_lt_iff_lt_one_right' [covariant_class α α (swap (*)) (<)] [contravariant_class α α (swap (*)) (<)] @@ -213,6 +320,9 @@ end has_lt section preorder variable [preorder α] +/-! Lemmas in the form of `b ≤ c → a ≤ 1 → b * a ≤ c`, +which assume left covariance. -/ + @[to_additive] lemma mul_le_of_le_of_le_one [covariant_class α α (*) (≤)] {a b c : α} (hbc : b ≤ c) (ha : a ≤ 1) : b * a ≤ c := @@ -224,19 +334,44 @@ alias mul_le_of_le_of_le_one ← mul_le_one' attribute [to_additive add_nonpos] mul_le_one' @[to_additive] -lemma lt_mul_of_lt_of_one_le [covariant_class α α (*) (≤)] - {a b c : α} (hbc : b < c) (ha : 1 ≤ a) : b < c * a := -calc b < c : hbc - ... = c * 1 : (mul_one c).symm - ... ≤ c * a : mul_le_mul_left' ha c +lemma mul_lt_of_le_of_lt_one [covariant_class α α (*) (<)] + {a b c : α} (hbc : b ≤ c) (ha : a < 1) : b * a < c := +calc b * a < b * 1 : mul_lt_mul_left' ha b + ... = b : mul_one b + ... ≤ c : hbc @[to_additive] lemma mul_lt_of_lt_of_le_one [covariant_class α α (*) (≤)] - {a b c : α} (hbc : b < c) (ha : a ≤ 1) : b * a < c := + {a b c : α} (hbc : b < c) (ha : a ≤ 1) : b * a < c := calc b * a ≤ b * 1 : mul_le_mul_left' ha b ... = b : mul_one b ... < c : hbc +@[to_additive] +lemma mul_lt_of_lt_of_lt_one [covariant_class α α (*) (<)] + {a b c : α} (hbc : b < c) (ha : a < 1) : b * a < c := +calc b * a < b * 1 : mul_lt_mul_left' ha b + ... = b : mul_one b + ... < c : hbc + +/-- Assumes left covariance. The lemma assuming right covariance is `right.mul_lt_one`. -/ +@[to_additive "Assumes left covariance. The lemma assuming right covariance is `right.add_neg`."] +lemma left.mul_lt_one [covariant_class α α (*) (<)] + {a b : α} (ha : a < 1) (hb : b < 1) : a * b < 1 := +calc a * b < a * 1 : mul_lt_mul_left' hb a + ... = a : mul_one a + ... < 1 : ha + +/-! Lemmas in the form of `b ≤ c → 1 ≤ a → b ≤ c * a`, +which assume left covariance. -/ + +@[to_additive] +lemma le_mul_of_le_of_le_one [covariant_class α α (*) (≤)] + {a b c : α} (hbc : b ≤ c) (ha : 1 ≤ a) : b ≤ c * a := +calc b ≤ c : hbc + ... = c * 1 : (mul_one c).symm + ... ≤ c * a : mul_le_mul_left' ha c + @[to_additive] lemma lt_mul_of_le_of_one_lt [covariant_class α α (*) (<)] {a b c : α} (hbc : b ≤ c) (ha : 1 < a) : b < c * a := @@ -245,34 +380,95 @@ calc b ≤ c : hbc ... < c * a : mul_lt_mul_left' ha c @[to_additive] -lemma mul_lt_of_le_one_of_lt [covariant_class α α (swap (*)) (≤)] - {a b c : α} (ha : a ≤ 1) (hb : b < c) : a * b < c := -calc a * b ≤ 1 * b : mul_le_mul_right' ha b - ... = b : one_mul b - ... < c : hb +lemma lt_mul_of_lt_of_one_le [covariant_class α α (*) (≤)] + {a b c : α} (hbc : b < c) (ha : 1 ≤ a) : b < c * a := +calc b < c : hbc + ... = c * 1 : (mul_one c).symm + ... ≤ c * a : mul_le_mul_left' ha c + +@[to_additive] +lemma lt_mul_of_lt_of_one_lt [covariant_class α α (*) (<)] + {a b c : α} (hbc : b < c) (ha : 1 < a) : b < c * a := +calc b < c : hbc + ... = c * 1 : (mul_one c).symm + ... < c * a : mul_lt_mul_left' ha c + +@[to_additive] +lemma lt_mul_of_lt_of_one_le' [covariant_class α α (*) (≤)] + {a b c : α} (hbc : b < c) (ha : 1 ≤ a) : b < c * a := +hbc.trans_le $ le_mul_of_one_le_right' ha + +@[to_additive] +lemma lt_mul_of_lt_of_one_lt' [covariant_class α α (*) (≤)] + {a b c : α} (hbc : b < c) (ha : 1 < a) : b < c * a := +lt_mul_of_lt_of_one_le' hbc ha.le + +@[to_additive] +lemma le_mul_of_le_of_one_le [covariant_class α α (*) (≤)] + {a b c : α} (hbc : b ≤ c) (ha : 1 ≤ a) : b ≤ c * a := +calc b ≤ c : hbc + ... = c * 1 : (mul_one c).symm + ... ≤ c * a : mul_le_mul_left' ha c + +@[to_additive add_nonneg] +lemma one_le_mul_right [covariant_class α α (*) (≤)] + {a b : α} (ha : 1 ≤ a) (hb : 1 ≤ b) : 1 ≤ a * b := +calc 1 ≤ a : ha + ... = a * 1 : (mul_one a).symm + ... ≤ a * b : mul_le_mul_left' hb a + +/- This lemma is present to mimick the name of an existing one. -/ +@[to_additive add_nonneg] +lemma one_le_mul [covariant_class α α (*) (≤)] + {a b : α} (ha : 1 ≤ a) (hb : 1 ≤ b) : 1 ≤ a * b := +le_mul_of_le_of_le_one ha hb + +@[to_additive add_pos_of_pos_of_nonneg] +lemma one_lt_mul_of_lt_of_le' [covariant_class α α (*) (≤)] + {a b : α} (ha : 1 < a) (hb : 1 ≤ b) : 1 < a * b := +lt_of_lt_of_le ha $ le_mul_of_one_le_right' hb + +@[to_additive add_pos] +lemma one_lt_mul' [covariant_class α α (*) (≤)] + {a b : α} (ha : 1 < a) (hb : 1 < b) : 1 < a * b := +one_lt_mul_of_lt_of_le' ha hb.le + +/-! Lemmas in the form of `a ≤ 1 → b ≤ c → a * b ≤ c`, +which assume right covariance. -/ @[to_additive] lemma mul_le_of_le_one_of_le [covariant_class α α (swap (*)) (≤)] - {a b c : α} (ha : a ≤ 1) (hbc : b ≤ c) : - a * b ≤ c := + {a b c : α} (ha : a ≤ 1) (hbc : b ≤ c) : a * b ≤ c := calc a * b ≤ 1 * b : mul_le_mul_right' ha b ... = b : one_mul b ... ≤ c : hbc @[to_additive] -lemma le_mul_of_one_le_of_le [covariant_class α α (swap (*)) (≤)] - {a b c: α} (ha : 1 ≤ a) (hbc : b ≤ c) : b ≤ a * c := -calc b ≤ c : hbc - ... = 1 * c : (one_mul c).symm - ... ≤ a * c : mul_le_mul_right' ha c +lemma mul_lt_of_lt_one_of_le [covariant_class α α (swap (*)) (<)] + {a b c : α} (ha : a < 1) (hbc : b ≤ c) : a * b < c := +calc a * b < 1 * b : mul_lt_mul_right' ha b + ... = b : one_mul b + ... ≤ c : hbc -/-- Assumes left covariance. The lemma assuming right covariance is `right.mul_lt_one`. -/ -@[to_additive "Assumes left covariance. The lemma assuming right covariance is `right.add_neg`."] -lemma left.mul_lt_one [covariant_class α α (*) (<)] - {a b : α} (ha : a < 1) (hb : b < 1) : a * b < 1 := -calc a * b < a * 1 : mul_lt_mul_left' hb a - ... = a : mul_one a - ... < 1 : ha +@[to_additive] +lemma mul_lt_of_le_one_of_lt [covariant_class α α (swap (*)) (≤)] + {a b c : α} (ha : a ≤ 1) (hb : b < c) : a * b < c := +calc a * b ≤ 1 * b : mul_le_mul_right' ha b + ... = b : one_mul b + ... < c : hb + +@[to_additive] +lemma mul_lt_of_lt_one_of_lt [covariant_class α α (swap (*)) (<)] + {a b c : α} (ha : a < 1) (hb : b < c) : a * b < c := +calc a * b < 1 * b : mul_lt_mul_right' ha b + ... = b : one_mul b + ... < c : hb + +@[to_additive add_le_of_nonpos_left] +lemma mul_le_of_le_one_left' [covariant_class α α (swap (*)) (≤)] + {a b : α} (h : b ≤ 1) : b * a ≤ a := +calc b * a ≤ 1 * a : mul_le_mul_right' h a + ... = a : one_mul a /-- Assumes right covariance. The lemma assuming left covariance is `left.mul_lt_one`. -/ @[to_additive "Assumes right covariance. The lemma assuming left covariance is `left.add_neg`"] @@ -282,20 +478,15 @@ calc a * b < 1 * b : mul_lt_mul_right' ha b ... = b : one_mul b ... < 1 : hb -@[to_additive] -lemma mul_lt_of_le_of_lt_one - [covariant_class α α (*) (<)] [covariant_class α α (swap (*)) (≤)] - {a b c: α} (hbc : b ≤ c) (ha : a < 1) : b * a < c := -calc b * a ≤ c * a : mul_le_mul_right' hbc a - ... < c * 1 : mul_lt_mul_left' ha c - ... = c : mul_one c +/-! Lemmas in the form of `1 ≤ a → b ≤ c → b ≤ a * c`, +which assume right covariance. -/ @[to_additive] -lemma mul_lt_of_lt_one_of_le [covariant_class α α (swap (*)) (<)] - {a b c : α} (ha : a < 1) (hbc : b ≤ c) : a * b < c := -calc a * b < 1 * b : mul_lt_mul_right' ha b - ... = b : one_mul b - ... ≤ c : hbc +lemma le_mul_of_one_le_of_le [covariant_class α α (swap (*)) (≤)] + {a b c : α} (ha : 1 ≤ a) (hbc : b ≤ c) : b ≤ a * c := +calc b ≤ c : hbc + ... = 1 * c : (one_mul c).symm + ... ≤ a * c : mul_le_mul_right' ha c @[to_additive] lemma lt_mul_of_one_lt_of_le [covariant_class α α (swap (*)) (<)] @@ -304,43 +495,29 @@ calc b ≤ c : hbc ... = 1 * c : (one_mul c).symm ... < a * c : mul_lt_mul_right' ha c -/-- Assumes left covariance. -/ -@[to_additive "Assumes left covariance."] -lemma le_mul_of_le_of_le_one [covariant_class α α (*) (≤)] - {a b c : α} (ha : c ≤ a) (hb : 1 ≤ b) : c ≤ a * b := -calc c ≤ a : ha - ... = a * 1 : (mul_one a).symm - ... ≤ a * b : mul_le_mul_left' hb a +@[to_additive] +lemma lt_mul_of_one_le_of_lt [covariant_class α α (swap (*)) (≤)] + {a b c : α} (ha : 1 ≤ a) (hbc : b < c) : b < a * c := +calc b < c : hbc + ... = 1 * c : (one_mul c).symm + ... ≤ a * c : mul_le_mul_right' ha c -/- This lemma is present to mimick the name of an existing one. -/ -@[to_additive add_nonneg] -lemma one_le_mul [covariant_class α α (*) (≤)] - {a b : α} (ha : 1 ≤ a) (hb : 1 ≤ b) : 1 ≤ a * b := -le_mul_of_le_of_le_one ha hb +@[to_additive] +lemma lt_mul_of_one_lt_of_lt [covariant_class α α (swap (*)) (≤)] + {a b c : α} (ha : 1 < a) (hbc : b < c) : b < a * c := +lt_mul_of_one_le_of_lt ha.le hbc -/-- Assumes left covariance. -/ -@[to_additive "Assumes left covariance."] -lemma lt_mul_of_lt_of_one_lt [covariant_class α α (*) (<)] - {a b c : α} (ha : c < a) (hb : 1 < b) : c < a * b := -calc c < a : ha - ... = a * 1 : (mul_one _).symm - ... < a * b : mul_lt_mul_left' hb a +@[to_additive le_add_of_nonneg_left] +lemma le_mul_of_one_le_left' [covariant_class α α (swap (*)) (≤)] + {a b : α} (h : 1 ≤ b) : a ≤ b * a := +calc a = 1 * a : (one_mul a).symm + ... ≤ b * a : mul_le_mul_right' h a -/-- Assumes left covariance. -/ -@[to_additive "Assumes left covariance."] -lemma mul_lt_of_lt_of_lt_one [covariant_class α α (*) (<)] - {a b c : α} (ha : a < c) (hb : b < 1) : a * b < c := -calc a * b < a * 1 : mul_lt_mul_left' hb a - ... = a : mul_one a - ... < c : ha - -/-- Assumes right covariance. -/ -@[to_additive "Assumes right covariance."] -lemma mul_lt_of_lt_one_of_lt [covariant_class α α (swap (*)) (<)] - {a b c : α} (ha : a < 1) (hb : b < c) : a * b < c := -calc a * b < 1 * b : mul_lt_mul_right' ha b - ... = b : one_mul b - ... < c : hb +@[to_additive lt_add_of_pos_left] +lemma lt_mul_of_one_lt_left' [covariant_class α α (swap (*)) (<)] + (a : α) {b : α} (h : 1 < b) : a < b * a := +calc a = 1 * a : (one_mul _).symm + ... < b * a : mul_lt_mul_right' h a /-- Assumes right covariance. -/ @[to_additive right.add_nonneg "Assumes right covariance."] @@ -350,41 +527,104 @@ calc 1 ≤ b : hb ... = 1 * b : (one_mul b).symm ... ≤ a * b : mul_le_mul_right' ha b +@[to_additive add_pos_of_nonneg_of_pos] +lemma one_lt_mul_of_le_of_lt' [covariant_class α α (swap (*)) (≤)] + {a b : α} (ha : 1 ≤ a) (hb : 1 < b) : 1 < a * b := +lt_of_lt_of_le hb $ le_mul_of_one_le_left' ha + /-- Assumes right covariance. -/ @[to_additive right.add_pos "Assumes right covariance."] lemma right.one_lt_mul [covariant_class α α (swap (*)) (<)] - {b : α} (hb : 1 < b) {a: α} (ha : 1 < a) : 1 < a * b := + {a b : α} (hb : 1 < b) (ha : 1 < a) : 1 < a * b := calc 1 < b : hb - ... = 1 * b : (one_mul _).symm + ... = 1 * b : (one_mul b).symm ... < a * b : mul_lt_mul_right' ha b -end preorder +@[to_additive] +lemma lt_of_mul_lt_of_one_le_left [covariant_class α α (*) (≤)] + {a b c : α} (h : a * b < c) (hle : 1 ≤ b) : a < c := +(le_mul_of_one_le_right' hle).trans_lt h +@[to_additive] +lemma le_of_mul_le_of_one_le_left [covariant_class α α (*) (≤)] + {a b c : α} (h : a * b ≤ c) (hle : 1 ≤ b) : a ≤ c := +(le_mul_of_one_le_right' hle).trans h -@[to_additive le_add_of_nonneg_right] -lemma le_mul_of_one_le_right' [has_le α] [covariant_class α α (*) (≤)] {a b : α} (h : 1 ≤ b) : - a ≤ a * b := -calc a = a * 1 : (mul_one _).symm - ... ≤ a * b : mul_le_mul_left' h a +@[to_additive] +lemma lt_of_lt_mul_of_le_one_left [covariant_class α α (*) (≤)] + {a b c : α} (h : a < b * c) (hle : c ≤ 1) : a < b := +h.trans_le (mul_le_of_le_one_right' hle) -@[to_additive add_le_of_nonpos_right] -lemma mul_le_of_le_one_right' [has_le α] [covariant_class α α (*) (≤)] {a b : α} (h : b ≤ 1) : - a * b ≤ a := -calc a * b ≤ a * 1 : mul_le_mul_left' h a - ... = a : mul_one a +@[to_additive] +lemma le_of_le_mul_of_le_one_left [covariant_class α α (*) (≤)] + {a b c : α} (h : a ≤ b * c) (hle : c ≤ 1) : a ≤ b := +h.trans (mul_le_of_le_one_right' hle) -end mul_one_class +@[to_additive] +lemma lt_of_mul_lt_of_one_le_right [covariant_class α α (swap (*)) (≤)] + {a b c : α} (h : a * b < c) (hle : 1 ≤ a) : b < c := +(le_mul_of_one_le_left' hle).trans_lt h @[to_additive] -lemma mul_left_cancel'' [has_mul α] [partial_order α] - [contravariant_class α α (*) (≤)] {a b c : α} (h : a * b = a * c) : b = c := -(le_of_mul_le_mul_left' h.le).antisymm (le_of_mul_le_mul_left' h.ge) +lemma le_of_mul_le_of_one_le_right [covariant_class α α (swap (*)) (≤)] + {a b c : α} (h : a * b ≤ c) (hle : 1 ≤ a) : b ≤ c := +(le_mul_of_one_le_left' hle).trans h @[to_additive] -lemma mul_right_cancel'' [has_mul α] [partial_order α] - [contravariant_class α α (swap (*)) (≤)] {a b c : α} (h : a * b = c * b) : - a = c := -le_antisymm (le_of_mul_le_mul_right' h.le) (le_of_mul_le_mul_right' h.ge) +lemma lt_of_lt_mul_of_le_one_right [covariant_class α α (swap (*)) (≤)] + {a b c : α} (h : a < b * c) (hle : b ≤ 1) : a < c := +h.trans_le (mul_le_of_le_one_left' hle) + +@[to_additive] +lemma le_of_le_mul_of_le_one_right [covariant_class α α (swap (*)) (≤)] + {a b c : α} (h : a ≤ b * c) (hle : b ≤ 1) : a ≤ c := +h.trans (mul_le_of_le_one_left' hle) + +end preorder + +section partial_order +variables [partial_order α] + +@[to_additive] +lemma mul_eq_one_iff' [covariant_class α α (*) (≤)] [covariant_class α α (swap (*)) (≤)] + {a b : α} (ha : 1 ≤ a) (hb : 1 ≤ b) : a * b = 1 ↔ a = 1 ∧ b = 1 := +iff.intro + (assume hab : a * b = 1, + have a ≤ 1, from hab ▸ le_mul_of_le_of_one_le le_rfl hb, + have a = 1, from le_antisymm this ha, + have b ≤ 1, from hab ▸ le_mul_of_one_le_of_le ha le_rfl, + have b = 1, from le_antisymm this hb, + and.intro ‹a = 1› ‹b = 1›) + (assume ⟨ha', hb'⟩, by rw [ha', hb', mul_one]) + +end partial_order + +section linear_order +variables [linear_order α] + +lemma exists_square_le [covariant_class α α (*) (<)] + (a : α) : ∃ (b : α), b * b ≤ a := +begin + by_cases h : a < 1, + { use a, + have : a*a < a*1, + exact mul_lt_mul_left' h a, + rw mul_one at this, + exact le_of_lt this }, + { use 1, + push_neg at h, + rwa mul_one } +end + +end linear_order + +end mul_one_class + +section semigroup +variables [semigroup α] + +section partial_order +variables [partial_order α] /- This is not instance, since we want to have an instance from `left_cancel_semigroup`s to the appropriate `covariant_class`. -/ @@ -393,7 +633,7 @@ to the appropriate `covariant_class`. -/ @[to_additive "An additive semigroup with a partial order and satisfying `left_cancel_add_semigroup` (i.e. `c + a < c + b → a < b`) is a `left_cancel add_semigroup`."] -def contravariant.to_left_cancel_semigroup [semigroup α] [partial_order α] +def contravariant.to_left_cancel_semigroup [contravariant_class α α (*) (≤)] : left_cancel_semigroup α := { mul_left_cancel := λ a b c, mul_left_cancel'' @@ -406,42 +646,13 @@ to the appropriate `covariant_class`. -/ @[to_additive "An additive semigroup with a partial order and satisfying `right_cancel_add_semigroup` (`a + c < b + c → a < b`) is a `right_cancel add_semigroup`."] -def contravariant.to_right_cancel_semigroup [semigroup α] [partial_order α] +def contravariant.to_right_cancel_semigroup [contravariant_class α α (swap (*)) (≤)] : right_cancel_semigroup α := { mul_right_cancel := λ a b c, mul_right_cancel'' ..‹semigroup α› } -variables {a b c d : α} - -section left -variables [preorder α] - -section has_mul -variables [has_mul α] - -@[to_additive] -lemma mul_lt_mul_of_lt_of_lt - [covariant_class α α (*) (<)] [covariant_class α α (swap (*)) (<)] - (h₁ : a < b) (h₂ : c < d) : a * c < b * d := -calc a * c < a * d : mul_lt_mul_left' h₂ a - ... < b * d : mul_lt_mul_right' h₁ d - -section contravariant_mul_lt_left_le_right -variables [covariant_class α α (*) (<)] [covariant_class α α (swap (*)) (≤)] - -@[to_additive] -lemma mul_lt_mul_of_le_of_lt - (h₁ : a ≤ b) (h₂ : c < d) : a * c < b * d := -(mul_le_mul_right' h₁ _).trans_lt (mul_lt_mul_left' h₂ b) - -@[to_additive add_lt_add] -lemma mul_lt_mul''' (h₁ : a < b) (h₂ : c < d) : a * c < b * d := -mul_lt_mul_of_le_of_lt h₁.le h₂ - -end contravariant_mul_lt_left_le_right - -@[to_additive] lemma mul_eq_mul_iff_eq_and_eq {α : Type*} [semigroup α] [partial_order α] +@[to_additive] lemma mul_eq_mul_iff_eq_and_eq [contravariant_class α α (*) (≤)] [covariant_class α α (swap (*)) (≤)] [covariant_class α α (*) (<)] [contravariant_class α α (swap (*)) (≤)] {a b c d : α} (hac : a ≤ c) (hbd : b ≤ d) : a * b = c * d ↔ a = c ∧ b = d := @@ -454,234 +665,10 @@ begin exact ((mul_lt_mul''' hac hbd).ne h).elim, end -variable [covariant_class α α (*) (≤)] - -@[to_additive] -lemma mul_lt_of_mul_lt_left (h : a * b < c) (hle : d ≤ b) : - a * d < c := -(mul_le_mul_left' hle a).trans_lt h - -@[to_additive] -lemma mul_le_of_mul_le_left (h : a * b ≤ c) (hle : d ≤ b) : - a * d ≤ c := -@act_rel_of_rel_of_act_rel _ _ _ (≤) _ ⟨λ _ _ _, le_trans⟩ a _ _ _ hle h - -@[to_additive] -lemma lt_mul_of_lt_mul_left (h : a < b * c) (hle : c ≤ d) : - a < b * d := -h.trans_le (mul_le_mul_left' hle b) - -@[to_additive] -lemma le_mul_of_le_mul_left (h : a ≤ b * c) (hle : c ≤ d) : - a ≤ b * d := -@rel_act_of_rel_of_rel_act _ _ _ (≤) _ ⟨λ _ _ _, le_trans⟩ b _ _ _ hle h - -@[to_additive] -lemma mul_lt_mul_of_lt_of_le [covariant_class α α (swap (*)) (<)] - (h₁ : a < b) (h₂ : c ≤ d) : a * c < b * d := -(mul_le_mul_left' h₂ _).trans_lt (mul_lt_mul_right' h₁ d) - -end has_mul - -/-! Here we start using properties of one, on the left. -/ -section mul_one_class -variables [mul_one_class α] [covariant_class α α (*) (≤)] - -@[to_additive] -lemma lt_of_mul_lt_of_one_le_left (h : a * b < c) (hle : 1 ≤ b) : a < c := -(le_mul_of_one_le_right' hle).trans_lt h - -@[to_additive] -lemma le_of_mul_le_of_one_le_left (h : a * b ≤ c) (hle : 1 ≤ b) : a ≤ c := -(le_mul_of_one_le_right' hle).trans h - -@[to_additive] -lemma lt_of_lt_mul_of_le_one_left (h : a < b * c) (hle : c ≤ 1) : a < b := -h.trans_le (mul_le_of_le_one_right' hle) - -@[to_additive] -lemma le_of_le_mul_of_le_one_left (h : a ≤ b * c) (hle : c ≤ 1) : a ≤ b := -h.trans (mul_le_of_le_one_right' hle) - -end mul_one_class - -end left - -section right - -section preorder -variables [preorder α] - -section has_mul -variables [has_mul α] - -variable [covariant_class α α (swap (*)) (≤)] - -@[to_additive] -lemma mul_lt_of_mul_lt_right (h : a * b < c) (hle : d ≤ a) : - d * b < c := -(mul_le_mul_right' hle b).trans_lt h - -@[to_additive] -lemma mul_le_of_mul_le_right (h : a * b ≤ c) (hle : d ≤ a) : - d * b ≤ c := -(mul_le_mul_right' hle b).trans h - -@[to_additive] -lemma lt_mul_of_lt_mul_right (h : a < b * c) (hle : b ≤ d) : - a < d * c := -h.trans_le (mul_le_mul_right' hle c) - -@[to_additive] -lemma le_mul_of_le_mul_right (h : a ≤ b * c) (hle : b ≤ d) : - a ≤ d * c := -h.trans (mul_le_mul_right' hle c) - -variable [covariant_class α α (*) (≤)] - -@[to_additive add_le_add] -lemma mul_le_mul' (h₁ : a ≤ b) (h₂ : c ≤ d) : a * c ≤ b * d := -(mul_le_mul_left' h₂ _).trans (mul_le_mul_right' h₁ d) - -@[to_additive] -lemma mul_le_mul_three {e f : α} (h₁ : a ≤ d) (h₂ : b ≤ e) (h₃ : c ≤ f) : - a * b * c ≤ d * e * f := -mul_le_mul' (mul_le_mul' h₁ h₂) h₃ - -end has_mul - -/-! Here we start using properties of one, on the right. -/ -section mul_one_class -variables [mul_one_class α] - -section le_right -variable [covariant_class α α (swap (*)) (≤)] - -@[to_additive le_add_of_nonneg_left] -lemma le_mul_of_one_le_left' (h : 1 ≤ b) : a ≤ b * a := -calc a = 1 * a : (one_mul a).symm - ... ≤ b * a : mul_le_mul_right' h a - -@[to_additive add_le_of_nonpos_left] -lemma mul_le_of_le_one_left' (h : b ≤ 1) : b * a ≤ a := -calc b * a ≤ 1 * a : mul_le_mul_right' h a - ... = a : one_mul a - -@[to_additive] -lemma lt_of_mul_lt_of_one_le_right (h : a * b < c) (hle : 1 ≤ a) : b < c := -(le_mul_of_one_le_left' hle).trans_lt h - -@[to_additive] -lemma le_of_mul_le_of_one_le_right (h : a * b ≤ c) (hle : 1 ≤ a) : b ≤ c := -(le_mul_of_one_le_left' hle).trans h - -@[to_additive] -lemma lt_of_lt_mul_of_le_one_right (h : a < b * c) (hle : b ≤ 1) : a < c := -h.trans_le (mul_le_of_le_one_left' hle) - -@[to_additive] -lemma le_of_le_mul_of_le_one_right (h : a ≤ b * c) (hle : b ≤ 1) : a ≤ c := -h.trans (mul_le_of_le_one_left' hle) - -end le_right - -section lt_right - -@[to_additive lt_add_of_pos_left] -lemma lt_mul_of_one_lt_left' [covariant_class α α (swap (*)) (<)] - (a : α) {b : α} (h : 1 < b) : a < b * a := -calc a = 1 * a : (one_mul _).symm - ... < b * a : mul_lt_mul_right' h a - -end lt_right - -end mul_one_class - -end preorder - -end right - -section preorder -variables [preorder α] - -section mul_one_class -variables [mul_one_class α] - -section covariant_left -variable [covariant_class α α (*) (≤)] - -@[to_additive add_pos_of_pos_of_nonneg] -lemma one_lt_mul_of_lt_of_le' (ha : 1 < a) (hb : 1 ≤ b) : 1 < a * b := -lt_of_lt_of_le ha $ le_mul_of_one_le_right' hb - -@[to_additive add_pos] -lemma one_lt_mul' (ha : 1 < a) (hb : 1 < b) : 1 < a * b := -one_lt_mul_of_lt_of_le' ha hb.le - -@[to_additive] -lemma lt_mul_of_lt_of_one_le' (hbc : b < c) (ha : 1 ≤ a) : - b < c * a := -hbc.trans_le $ le_mul_of_one_le_right' ha - -@[to_additive] -lemma lt_mul_of_lt_of_one_lt' (hbc : b < c) (ha : 1 < a) : - b < c * a := -lt_mul_of_lt_of_one_le' hbc ha.le - -@[to_additive] -lemma le_mul_of_le_of_one_le (hbc : b ≤ c) (ha : 1 ≤ a) : b ≤ c * a := -calc b ≤ c : hbc - ... = c * 1 : (mul_one c).symm - ... ≤ c * a : mul_le_mul_left' ha c - -@[to_additive add_nonneg] -lemma one_le_mul_right (ha : 1 ≤ a) (hb : 1 ≤ b) : 1 ≤ a * b := -calc 1 ≤ a : ha - ... = a * 1 : (mul_one a).symm - ... ≤ a * b : mul_le_mul_left' hb a - -end covariant_left - -section covariant_right -variable [covariant_class α α (swap (*)) (≤)] - -@[to_additive add_pos_of_nonneg_of_pos] -lemma one_lt_mul_of_le_of_lt' (ha : 1 ≤ a) (hb : 1 < b) : 1 < a * b := -lt_of_lt_of_le hb $ le_mul_of_one_le_left' ha - -@[to_additive] -lemma lt_mul_of_one_le_of_lt (ha : 1 ≤ a) (hbc : b < c) : b < a * c := -hbc.trans_le $ le_mul_of_one_le_left' ha - -@[to_additive] -lemma lt_mul_of_one_lt_of_lt (ha : 1 < a) (hbc : b < c) : b < a * c := -lt_mul_of_one_le_of_lt ha.le hbc - -end covariant_right - -end mul_one_class - -end preorder - -section partial_order - -/-! Properties assuming `partial_order`. -/ -variables [mul_one_class α] [partial_order α] - [covariant_class α α (*) (≤)] [covariant_class α α (swap (*)) (≤)] - -@[to_additive] -lemma mul_eq_one_iff' (ha : 1 ≤ a) (hb : 1 ≤ b) : a * b = 1 ↔ a = 1 ∧ b = 1 := -iff.intro - (assume hab : a * b = 1, - have a ≤ 1, from hab ▸ le_mul_of_le_of_one_le le_rfl hb, - have a = 1, from le_antisymm this ha, - have b ≤ 1, from hab ▸ le_mul_of_one_le_of_le ha le_rfl, - have b = 1, from le_antisymm this hb, - and.intro ‹a = 1› ‹b = 1›) - (assume ⟨ha', hb'⟩, by rw [ha', hb', mul_one]) - end partial_order +end semigroup + section mono variables [has_mul α] [preorder α] [preorder β] {f g : β → α} From 559f58b11032054f3ee8ee7099887558350a5dbd Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 7 May 2022 07:07:16 +0000 Subject: [PATCH 06/37] feat(order/filter): add a few lemmas (#13985) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * weaken assumptions of `filter.has_antitone_basis.tendsto` from `[semilattice_sup ι] [nonempty ι]` to `[preorder ι]`; * add `filter.has_antitone_basis.tendsto`, `filter.has_antitone_basis.mem`, `filter.has_antitone_basis.tendsto_small_sets`. --- src/order/filter/at_top_bot.lean | 11 ++++++++--- src/order/filter/bases.lean | 4 ++++ src/order/filter/small_sets.lean | 5 +++++ 3 files changed, 17 insertions(+), 3 deletions(-) diff --git a/src/order/filter/at_top_bot.lean b/src/order/filter/at_top_bot.lean index f28d73a903a54..47bd815c63aa1 100644 --- a/src/order/filter/at_top_bot.lean +++ b/src/order/filter/at_top_bot.lean @@ -1288,11 +1288,16 @@ by rw [map_at_top_eq, map_at_top_eq]; from (le_infi $ assume b, let ⟨v, hv⟩ := h_eq b in infi_le_of_le v $ by simp [set.image_subset_iff]; exact hv) -lemma has_antitone_basis.tendsto [semilattice_sup ι] [nonempty ι] {l : filter α} +lemma has_antitone_basis.eventually_subset [preorder ι] {l : filter α} + {s : ι → set α} (hl : l.has_antitone_basis s) {t : set α} (ht : t ∈ l) : + ∀ᶠ i in at_top, s i ⊆ t := +let ⟨i, _, hi⟩ := hl.to_has_basis.mem_iff.1 ht +in (eventually_ge_at_top i).mono $ λ j hj, (hl.antitone hj).trans hi + +protected lemma has_antitone_basis.tendsto [preorder ι] {l : filter α} {s : ι → set α} (hl : l.has_antitone_basis s) {φ : ι → α} (h : ∀ i : ι, φ i ∈ s i) : tendsto φ at_top l := -(at_top_basis.tendsto_iff hl.to_has_basis).2 $ assume i hi, - ⟨i, trivial, λ j hij, hl.antitone hij (h _)⟩ +λ t ht, mem_map.2 $ (hl.eventually_subset ht).mono $ λ i hi, hi (h i) /-- If `f` is a nontrivial countably generated filter, then there exists a sequence that converges to `f`. -/ diff --git a/src/order/filter/bases.lean b/src/order/filter/bases.lean index ad29dd10e4acc..486c8b7415f61 100644 --- a/src/order/filter/bases.lean +++ b/src/order/filter/bases.lean @@ -803,6 +803,10 @@ countable_binfi_eq_infi_seq' Bcbl 𝓟 principal_univ section is_countably_generated +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 + /-- 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/order/filter/small_sets.lean b/src/order/filter/small_sets.lean index 50e6fc17bf75a..e47f8f397b6b1 100644 --- a/src/order/filter/small_sets.lean +++ b/src/order/filter/small_sets.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot, Floris van Doorn, Yury Kudryashov -/ import order.filter.lift +import order.filter.at_top_bot /-! # The filter of small sets @@ -55,6 +56,10 @@ lemma eventually_small_sets' {p : set α → Prop} (hp : ∀ ⦃s t⦄, s ⊆ t eventually_small_sets.trans $ exists₂_congr $ λ s hsf, ⟨λ H, H s (subset.refl s), λ hs t ht, hp ht hs⟩ +lemma has_antitone_basis.tendsto_small_sets {ι} [preorder ι] {s : ι → set α} + (hl : l.has_antitone_basis s) : tendsto s at_top l.small_sets := +tendsto_small_sets_iff.2 $ λ t ht, hl.eventually_subset ht + @[mono] lemma monotone_small_sets : monotone (@small_sets α) := monotone_lift' monotone_id monotone_const From f8bc097de20ed70d1d6ad4f9cd65a9ff89892948 Mon Sep 17 00:00:00 2001 From: Haruhisa Enomoto <73296681+haruhisa-enomoto@users.noreply.github.com> Date: Sat, 7 May 2022 08:20:10 +0000 Subject: [PATCH 07/37] =?UTF-8?q?feat(algebra/module/linear=5Fmap):=20`R?= =?UTF-8?q?=E1=B5=90=E1=B5=92=E1=B5=96`=20is=20isomorphic=20to=20`module.E?= =?UTF-8?q?nd=20R=20R`=20(#13931)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds the canonical (semi)ring isomorphism from `Rᵐᵒᵖ` to `module.End R R` for a (semi)ring `R`, given by the right multiplication. --- src/algebra/module/linear_map.lean | 23 +++++++++++++++++++++++ 1 file changed, 23 insertions(+) diff --git a/src/algebra/module/linear_map.lean b/src/algebra/module/linear_map.lean index 74d51e850b9d1..abda3d037225b 100644 --- a/src/algebra/module/linear_map.lean +++ b/src/algebra/module/linear_map.lean @@ -314,6 +314,9 @@ ext $ λ x, by rw [← mul_one x, ← smul_eq_mul, f.map_smulₛₗ, g.map_smul theorem ext_ring_iff {σ : R →+* R} {f g : R →ₛₗ[σ] M} : f = g ↔ f 1 = g 1 := ⟨λ h, h ▸ rfl, ext_ring⟩ +@[ext] theorem ext_ring_op {σ : Rᵐᵒᵖ →+* S} {f g : R →ₛₗ[σ] M₃} (h : f 1 = g 1) : f = g := +ext $ λ x, by rw [← one_mul x, ← op_smul_eq_mul, f.map_smulₛₗ, g.map_smulₛₗ, h] + end /-- Interpret a `ring_hom` `f` as an `f`-semilinear map. -/ @@ -843,4 +846,24 @@ def to_module_End : S →+* module.End R M := map_add' := λ f g, linear_map.ext $ add_smul _ _, ..distrib_mul_action.to_module_End R M } +/-- The canonical (semi)ring isomorphism from `Rᵐᵒᵖ` to `module.End R R` induced by the right +multiplication. -/ +@[simps] +def module_End_self : Rᵐᵒᵖ ≃+* module.End R R := +{ to_fun := distrib_mul_action.to_linear_map R R, + inv_fun := λ f, mul_opposite.op (f 1), + left_inv := mul_one, + right_inv := λ f, linear_map.ext_ring $ one_mul _, + ..module.to_module_End R R } + +/-- The canonical (semi)ring isomorphism from `R` to `module.End Rᵐᵒᵖ R` induced by the left +multiplication. -/ +@[simps] +def module_End_self_op : R ≃+* module.End Rᵐᵒᵖ R := +{ to_fun := distrib_mul_action.to_linear_map _ _, + inv_fun := λ f, f 1, + left_inv := mul_one, + right_inv := λ f, linear_map.ext_ring_op $ mul_one _, + ..module.to_module_End _ _ } + end module From b2aa27e319a83aa45e880a58f332278f09ebdfa4 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Sat, 7 May 2022 09:57:51 +0000 Subject: [PATCH 08/37] feat(analysis/calculus/deriv): generalize some lemmas (#13575) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The types of scalar and codomain can be different now. For example, these lemmas can be used for `f : ℝ → ℂ` `f' : ℝ →L[ℝ] ℂ` now. --- src/analysis/calculus/deriv.lean | 55 +++++++++++++++++--------------- 1 file changed, 30 insertions(+), 25 deletions(-) diff --git a/src/analysis/calculus/deriv.lean b/src/analysis/calculus/deriv.lean index 438d7efa507c8..be36dec44fa9d 100644 --- a/src/analysis/calculus/deriv.lean +++ b/src/analysis/calculus/deriv.lean @@ -1109,32 +1109,37 @@ lemma deriv.scomp /-! ### Derivative of the composition of a scalar and vector functions -/ -theorem has_deriv_at_filter.comp_has_fderiv_at_filter {f : E → 𝕜} {f' : E →L[𝕜] 𝕜} (x) - {L' : filter E} (hh₁ : has_deriv_at_filter h₁ h₁' (f x) L) - (hf : has_fderiv_at_filter f f' x L') (hL : tendsto f L' L) : - has_fderiv_at_filter (h₁ ∘ f) (h₁' • f') x L' := -by { convert hh₁.comp x hf hL, ext x, simp [mul_comm] } - -theorem has_strict_deriv_at.comp_has_strict_fderiv_at {f : E → 𝕜} {f' : E →L[𝕜] 𝕜} (x) - (hh₁ : has_strict_deriv_at h₁ h₁' (f x)) (hf : has_strict_fderiv_at f f' x) : - has_strict_fderiv_at (h₁ ∘ f) (h₁' • f') x := -by { rw has_strict_deriv_at at hh₁, convert hh₁.comp x hf, ext x, simp [mul_comm] } - -theorem has_deriv_at.comp_has_fderiv_at {f : E → 𝕜} {f' : E →L[𝕜] 𝕜} (x) - (hh₁ : has_deriv_at h₁ h₁' (f x)) (hf : has_fderiv_at f f' x) : - has_fderiv_at (h₁ ∘ f) (h₁' • f') x := -hh₁.comp_has_fderiv_at_filter x hf hf.continuous_at - -theorem has_deriv_at.comp_has_fderiv_within_at {f : E → 𝕜} {f' : E →L[𝕜] 𝕜} {s} (x) - (hh₁ : has_deriv_at h₁ h₁' (f x)) (hf : has_fderiv_within_at f f' s x) : - has_fderiv_within_at (h₁ ∘ f) (h₁' • f') s x := -hh₁.comp_has_fderiv_at_filter x hf hf.continuous_within_at - -theorem has_deriv_within_at.comp_has_fderiv_within_at {f : E → 𝕜} {f' : E →L[𝕜] 𝕜} {s t} (x) - (hh₁ : has_deriv_within_at h₁ h₁' t (f x)) (hf : has_fderiv_within_at f f' s x) +theorem has_deriv_at_filter.comp_has_fderiv_at_filter {f : E → 𝕜'} {f' : E →L[𝕜] 𝕜'} (x) + {L'' : filter E} (hh₂ : has_deriv_at_filter h₂ h₂' (f x) L') + (hf : has_fderiv_at_filter f f' x L'') (hL : tendsto f L'' L') : + has_fderiv_at_filter (h₂ ∘ f) (h₂' • f') x L'' := +by { convert (hh₂.restrict_scalars 𝕜).comp x hf hL, ext x, simp [mul_comm] } + +theorem has_strict_deriv_at.comp_has_strict_fderiv_at {f : E → 𝕜'} {f' : E →L[𝕜] 𝕜'} (x) + (hh : has_strict_deriv_at h₂ h₂' (f x)) (hf : has_strict_fderiv_at f f' x) : + has_strict_fderiv_at (h₂ ∘ f) (h₂' • f') x := +begin + rw has_strict_deriv_at at hh, + convert (hh.restrict_scalars 𝕜).comp x hf, + ext x, + simp [mul_comm] +end + +theorem has_deriv_at.comp_has_fderiv_at {f : E → 𝕜'} {f' : E →L[𝕜] 𝕜'} (x) + (hh : has_deriv_at h₂ h₂' (f x)) (hf : has_fderiv_at f f' x) : + has_fderiv_at (h₂ ∘ f) (h₂' • f') x := +hh.comp_has_fderiv_at_filter x hf hf.continuous_at + +theorem has_deriv_at.comp_has_fderiv_within_at {f : E → 𝕜'} {f' : E →L[𝕜] 𝕜'} {s} (x) + (hh : has_deriv_at h₂ h₂' (f x)) (hf : has_fderiv_within_at f f' s x) : + has_fderiv_within_at (h₂ ∘ f) (h₂' • f') s x := +hh.comp_has_fderiv_at_filter x hf hf.continuous_within_at + +theorem has_deriv_within_at.comp_has_fderiv_within_at {f : E → 𝕜'} {f' : E →L[𝕜] 𝕜'} {s t} (x) + (hh : has_deriv_within_at h₂ h₂' t (f x)) (hf : has_fderiv_within_at f f' s x) (hst : maps_to f s t) : - has_fderiv_within_at (h₁ ∘ f) (h₁' • f') s x := -hh₁.comp_has_fderiv_at_filter x hf $ hf.continuous_within_at.tendsto_nhds_within hst + has_fderiv_within_at (h₂ ∘ f) (h₂' • f') s x := +hh.comp_has_fderiv_at_filter x hf $ hf.continuous_within_at.tendsto_nhds_within hst /-! ### Derivative of the composition of two scalar functions -/ From 9134a8e763a4955bcb7ab64fa2c43b13552f6d8e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 7 May 2022 09:57:53 +0000 Subject: [PATCH 09/37] feat(combinatorics/simple_graph/hasse): Hasse diagram and path graph (#13959) Define the Hasse diagram of an order and the path graph on `n` vertices as the Hasse diagram of `fin n`. --- src/combinatorics/simple_graph/basic.lean | 8 +- .../simple_graph/connectivity.lean | 5 +- src/combinatorics/simple_graph/hasse.lean | 75 +++++++++++++++++++ 3 files changed, 82 insertions(+), 6 deletions(-) create mode 100644 src/combinatorics/simple_graph/hasse.lean diff --git a/src/combinatorics/simple_graph/basic.lean b/src/combinatorics/simple_graph/basic.lean index 1216aff11901e..9934190429130 100644 --- a/src/combinatorics/simple_graph/basic.lean +++ b/src/combinatorics/simple_graph/basic.lean @@ -88,6 +88,9 @@ structure simple_graph (V : Type u) := (symm : symmetric adj . obviously) (loopless : irreflexive adj . obviously) +noncomputable instance {V : Type u} [fintype V] : fintype (simple_graph V) := +by { classical, exact fintype.of_injective simple_graph.adj simple_graph.ext } + /-- Construct the simple graph induced by the given relation. It symmetrizes the relation and makes it irreflexive. @@ -97,9 +100,6 @@ def simple_graph.from_rel {V : Type u} (r : V → V → Prop) : simple_graph V : symm := λ a b ⟨hn, hr⟩, ⟨hn.symm, hr.symm⟩, loopless := λ a ⟨hn, _⟩, hn rfl } -noncomputable instance {V : Type u} [fintype V] : fintype (simple_graph V) := -by { classical, exact fintype.of_injective simple_graph.adj simple_graph.ext } - @[simp] lemma simple_graph.from_rel_adj {V : Type u} (r : V → V → Prop) (v w : V) : (simple_graph.from_rel r).adj v w ↔ v ≠ w ∧ (r v w ∨ r w v) := @@ -137,7 +137,7 @@ 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} -@[simp] lemma irrefl {v : V} : ¬G.adj v v := G.loopless v +@[simp] protected lemma irrefl {v : V} : ¬G.adj v v := G.loopless v lemma adj_comm (u v : V) : G.adj u v ↔ G.adj v u := ⟨λ x, G.symm x, λ x, G.symm x⟩ diff --git a/src/combinatorics/simple_graph/connectivity.lean b/src/combinatorics/simple_graph/connectivity.lean index 32409689aa3dc..2a5c053bedf40 100644 --- a/src/combinatorics/simple_graph/connectivity.lean +++ b/src/combinatorics/simple_graph/connectivity.lean @@ -942,7 +942,8 @@ begin exact h.elim (λ q, hp q.to_path), end -@[refl] protected lemma reachable.refl {u : V} : G.reachable u u := by { fsplit, refl } +@[refl] protected lemma reachable.refl (u : V) : G.reachable u u := by { fsplit, refl } +protected lemma reachable.rfl {u : V} : G.reachable u u := reachable.refl _ @[symm] protected lemma reachable.symm {u v : V} (huv : G.reachable u v) : G.reachable v u := huv.elim (λ p, ⟨p.reverse⟩) @@ -986,7 +987,7 @@ of `h.preconnected u v`. -/ @[protect_proj] structure connected : Prop := (preconnected : G.preconnected) -(nonempty : nonempty V) +[nonempty : nonempty V] instance : has_coe_to_fun G.connected (λ _, Π (u v : V), G.reachable u v) := ⟨λ h, h.preconnected⟩ diff --git a/src/combinatorics/simple_graph/hasse.lean b/src/combinatorics/simple_graph/hasse.lean new file mode 100644 index 0000000000000..937d1e815c118 --- /dev/null +++ b/src/combinatorics/simple_graph/hasse.lean @@ -0,0 +1,75 @@ +/- +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 combinatorics.simple_graph.connectivity +import data.fin.succ_pred +import order.succ_pred.relation + +/-! +# The Hasse diagram as a graph + +This file defines the Hasse diagram of an order (graph of `covby`, the covering relation) and the +path graph on `n` vertices. + +## Main declarations + +* `simple_graph.hasse`: Hasse diagram of an order. +* `simple_graph.path_graph`: Path graph on `n` vertices. +-/ + +open order order_dual relation + +namespace simple_graph +variables (α β : Type*) + +section preorder +variables [preorder α] [preorder β] + +/-- The Hasse diagram of an order as a simple graph. The graph of the covering relation. -/ +def hasse : simple_graph α := +{ adj := λ a b, a ⋖ b ∨ b ⋖ a, + symm := λ a b, or.symm, + loopless := λ a h, h.elim (irrefl _) (irrefl _) } + +variables {α β} {a b : α} + +@[simp] lemma hasse_adj : (hasse α).adj a b ↔ a ⋖ b ∨ b ⋖ a := iff.rfl + +/-- `αᵒᵈ` and `α` have the same Hasse diagram. -/ +def hasse_dual_iso : hasse αᵒᵈ ≃g hasse α := +{ map_rel_iff' := λ a b, by simp [or_comm], + ..of_dual } + +@[simp] lemma hasse_dual_iso_apply (a : αᵒᵈ) : hasse_dual_iso a = of_dual a := rfl +@[simp] lemma hasse_dual_iso_symm_apply (a : α) : hasse_dual_iso.symm a = to_dual a := rfl + +end preorder + +section linear_order +variables [linear_order α] + +lemma hasse_preconnected_of_succ [succ_order α] [is_succ_archimedean α] : (hasse α).preconnected := +λ a b, begin + rw reachable_iff_refl_trans_gen, + exact refl_trans_gen_of_succ _ (λ c hc, or.inl $ covby_succ_of_not_is_max hc.2.not_is_max) + (λ c hc, or.inr $ covby_succ_of_not_is_max hc.2.not_is_max), +end + +lemma hasse_preconnected_of_pred [pred_order α] [is_pred_archimedean α] : (hasse α).preconnected := +λ a b, begin + rw [reachable_iff_refl_trans_gen, ←refl_trans_gen_swap], + exact refl_trans_gen_of_pred _ (λ c hc, or.inl $ pred_covby_of_not_is_min hc.1.not_is_min) + (λ c hc, or.inr $ pred_covby_of_not_is_min hc.1.not_is_min), +end + +end linear_order + +/-- The path graph on `n` vertices. -/ +def path_graph (n : ℕ) : simple_graph (fin n) := hasse _ + +lemma path_graph_preconnected (n : ℕ) : (path_graph n).preconnected := hasse_preconnected_of_succ _ +lemma path_graph_connected (n : ℕ) : (path_graph (n + 1)).connected := ⟨path_graph_preconnected _⟩ + +end simple_graph From c24762232e82281e047f5b012fd7c33440d0be36 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Sat, 7 May 2022 09:57:54 +0000 Subject: [PATCH 10/37] feat(group_theory/group_action/units): simp lemma for scalar action of `is_unit.unit h` (#14006) --- src/algebra/group_with_zero/basic.lean | 5 +++++ src/group_theory/group_action/units.lean | 4 ++++ src/ring_theory/witt_vector/isocrystal.lean | 8 +++----- 3 files changed, 12 insertions(+), 5 deletions(-) diff --git a/src/algebra/group_with_zero/basic.lean b/src/algebra/group_with_zero/basic.lean index 6ef81c9ae4007..ac701b3712324 100644 --- a/src/algebra/group_with_zero/basic.lean +++ b/src/algebra/group_with_zero/basic.lean @@ -7,6 +7,7 @@ import algebra.group.inj_surj import algebra.group_with_zero.defs import algebra.hom.units import logic.nontrivial +import group_theory.group_action.units /-! # Groups with an adjoined zero element @@ -722,6 +723,10 @@ begin simpa only [eq_comm] using units.exists_iff_ne_zero.mpr h } end +@[simp] lemma smul_mk0 {α : Type*} [has_scalar G₀ α] {g : G₀} (hg : g ≠ 0) (a : α) : + (mk0 g hg) • a = g • a := +rfl + end units section group_with_zero diff --git a/src/group_theory/group_action/units.lean b/src/group_theory/group_action/units.lean index c7d3df43e723d..b21c464632f64 100644 --- a/src/group_theory/group_action/units.lean +++ b/src/group_theory/group_action/units.lean @@ -32,6 +32,10 @@ instance [monoid M] [has_scalar M α] : has_scalar Mˣ α := lemma smul_def [monoid M] [has_scalar M α] (m : Mˣ) (a : α) : m • a = (m : M) • a := rfl +@[simp] lemma smul_is_unit [monoid M] [has_scalar M α] {m : M} (hm : is_unit m) (a : α) : + hm.unit • a = m • a := +rfl + lemma _root_.is_unit.inv_smul [monoid α] {a : α} (h : is_unit a) : (h.unit)⁻¹ • a = 1 := h.coe_inv_mul diff --git a/src/ring_theory/witt_vector/isocrystal.lean b/src/ring_theory/witt_vector/isocrystal.lean index 1dce9a1916f94..a64ba515eb81f 100644 --- a/src/ring_theory/witt_vector/isocrystal.lean +++ b/src/ring_theory/witt_vector/isocrystal.lean @@ -192,11 +192,9 @@ begin intros c, rw [linear_equiv.trans_apply, linear_equiv.trans_apply, linear_equiv.smul_of_ne_zero_apply, linear_equiv.smul_of_ne_zero_apply, - show F (units.mk0 b hb • Φ(p,k) c) = _, from linear_equiv.map_smul _ _ _, - show F (units.mk0 b hb • c) = _, from linear_equiv.map_smul _ _ _], - simp only [hax, units.coe_mk0, linear_equiv.of_bijective_apply, - linear_map.to_span_singleton_apply, linear_equiv.map_smulₛₗ, - standard_one_dim_isocrystal.frobenius_apply, algebra.id.smul_eq_mul], + linear_equiv.map_smul, linear_equiv.map_smul], + simp only [hax, linear_equiv.of_bijective_apply, linear_map.to_span_singleton_apply, + linear_equiv.map_smulₛₗ, standard_one_dim_isocrystal.frobenius_apply, algebra.id.smul_eq_mul], simp only [←mul_smul], congr' 1, linear_combination (hmb, φ(p,k) c), From cf65daf0e2ff9730a15c401dfe8b77120c9a255c Mon Sep 17 00:00:00 2001 From: sgouezel Date: Sat, 7 May 2022 11:59:00 +0000 Subject: [PATCH 11/37] feat(probability/variance): define the variance of a random variable, prove its basic properties (#13912) --- docs/undergrad.yaml | 2 +- src/probability/integration.lean | 7 +- src/probability/variance.lean | 238 +++++++++++++++++++++++++++++++ 3 files changed, 245 insertions(+), 2 deletions(-) create mode 100644 src/probability/variance.lean diff --git a/docs/undergrad.yaml b/docs/undergrad.yaml index 4f27043f36f9e..4c527cb63e9bd 100644 --- a/docs/undergrad.yaml +++ b/docs/undergrad.yaml @@ -565,7 +565,7 @@ Probability Theory: $\mathrm{L}^p$ convergence: 'measure_theory.Lp' almost surely convergence: 'measure_theory.measure.ae' Markov inequality: 'measure_theory.mul_meas_ge_le_lintegral' - Tchebychev inequality: '' + Chebychev inequality: 'probability_theory.meas_ge_le_variance_div_sq' Levy's theorem: '' weak law of large numbers: '' strong law of large numbers: '' diff --git a/src/probability/integration.lean b/src/probability/integration.lean index 0fdd1e0fc7d43..25f55ffae3231 100644 --- a/src/probability/integration.lean +++ b/src/probability/integration.lean @@ -131,7 +131,7 @@ begin end /-- The product of two independent, integrable, real_valued random variables is integrable. -/ -lemma indep_fun.integrable_mul {β : Type*} {mβ : measurable_space β} {X Y : α → β} +lemma indep_fun.integrable_mul {β : Type*} [measurable_space β] {X Y : α → β} [normed_division_ring β] [borel_space β] (hXY : indep_fun X Y μ) (hX : integrable X μ) (hY : integrable Y μ) : integrable (X * Y) μ := @@ -235,6 +235,11 @@ 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 + /-- 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 `ℝ` satisfying appropriate integrability conditions. -/ diff --git a/src/probability/variance.lean b/src/probability/variance.lean new file mode 100644 index 0000000000000..804e9afef5a6a --- /dev/null +++ b/src/probability/variance.lean @@ -0,0 +1,238 @@ +/- +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 probability.notation +import probability.integration + +/-! +# Variance of random variables + +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., + `ℙ {ω | 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. +-/ + +open measure_theory filter finset + +noncomputable theory + +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 - (λ x, μ[f])) ^ 2] + +@[simp] lemma variance_zero {Ω : Type*} {m : measurable_space Ω} (μ : measure Ω) : + variance 0 μ = 0 := +by simp [variance] + +lemma variance_nonneg {Ω : Type*} {m : measurable_space Ω} (f : Ω → ℝ) (μ : measure Ω) : + 0 ≤ variance f μ := +integral_nonneg (λ x, sq_nonneg _) + +lemma variance_mul {Ω : Type*} {m : measurable_space Ω} (c : ℝ) (f : Ω → ℝ) (μ : measure Ω) : + variance (λ x, c * f x) μ = c^2 * variance f μ := +calc +variance (λ x, c * f x) μ + = ∫ 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 variance_smul {Ω : Type*} {m : measurable_space Ω} (c : ℝ) (f : Ω → ℝ) (μ : measure Ω) : + variance (c • f) μ = c^2 * variance f μ := +variance_mul c f μ + +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 μ := +begin + convert variance_smul (algebra_map A ℝ c) f μ, + { ext1 x, simp only [algebra_map_smul], }, + { simp only [algebra.smul_def, map_pow], } +end + +localized +"notation `Var[` X `]` := probability_theory.variance X measure_theory.measure_space.volume" +in probability_theory + +variables {Ω : Type*} [measure_space Ω] [is_probability_measure (volume : measure Ω)] + +lemma variance_def' {X : Ω → ℝ} (hX : mem_ℒp X 2) : + Var[X] = 𝔼[X^2] - 𝔼[X]^2 := +begin + rw [variance, sub_sq', integral_sub', integral_add'], rotate, + { exact hX.integrable_sq }, + { convert integrable_const (𝔼[X] ^ 2), + apply_instance }, + { apply hX.integrable_sq.add, + convert integrable_const (𝔼[X] ^ 2), + apply_instance }, + { exact ((hX.integrable ennreal.one_le_two).const_mul 2).mul_const' _ }, + simp only [integral_mul_right, pi.pow_apply, pi.mul_apply, pi.bit0_apply, pi.one_apply, + integral_const (integral ℙ X ^ 2), integral_mul_left (2 : ℝ), one_mul, + variance, pi.pow_apply, measure_univ, ennreal.one_to_real, algebra.id.smul_eq_mul], + ring, +end + +lemma variance_le_expectation_sq {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], + { exact integral_nonneg (λ a, sq_nonneg _) }, + { assume h, + have A : mem_ℒp (X - λ (x : Ω), 𝔼[X]) 2 ℙ := (mem_ℒp_two_iff_integrable_sq + (h_int.ae_strongly_measurable.sub ae_strongly_measurable_const)).2 h, + have B : mem_ℒp (λ (x : Ω), 𝔼[X]) 2 ℙ := mem_ℒp_const _, + apply hX, + convert A.add B, + simp } } +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) := +begin + have A : (ennreal.of_real c : ℝ≥0∞) ≠ 0, + by simp only [hc, ne.def, ennreal.of_real_eq_zero, not_le], + 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, + { 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₀] } +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) : + 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 : + by simp [variance_def' (hX.add hY), add_sq'] +... = (𝔼[X^2] + 𝔼[Y^2] + 2 * 𝔼[X * Y]) - (𝔼[X] + 𝔼[Y])^2 : +begin + simp only [pi.add_apply, pi.pow_apply, pi.mul_apply, mul_assoc], + rw [integral_add, integral_add, integral_add, integral_mul_left], + { exact hX.integrable ennreal.one_le_two }, + { exact hY.integrable ennreal.one_le_two }, + { exact hX.integrable_sq }, + { exact hY.integrable_sq }, + { exact hX.integrable_sq.add hY.integrable_sq }, + { apply integrable.const_mul, + exact h.integrable_mul (hX.integrable ennreal.one_le_two) (hY.integrable ennreal.one_le_two) } +end +... = (𝔼[X^2] + 𝔼[Y^2] + 2 * (𝔼[X] * 𝔼[Y])) - (𝔼[X] + 𝔼[Y])^2 : +begin + congr, + exact h.integral_mul_of_integrable + (hX.integrable ennreal.one_le_two) (hY.integrable ennreal.one_le_two), +end +... = Var[X] + Var[Y] : + by { simp only [variance_def', hX, hY, pi.pow_apply], ring } + +/-- 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 ι} + (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 + classical, + induction s using finset.induction_on with k s ks IH, + { simp only [finset.sum_empty, variance_zero] }, + rw [variance_def' (mem_ℒp_finset_sum' _ hs), sum_insert ks, sum_insert ks], + simp only [add_sq'], + calc 𝔼[X k ^ 2 + (∑ i in s, X i) ^ 2 + 2 * X k * ∑ i in s, X i] - 𝔼[X k + ∑ i in s, X i] ^ 2 + = (𝔼[X k ^ 2] + 𝔼[(∑ i in s, X i) ^ 2] + 𝔼[2 * X k * ∑ i in s, X i]) + - (𝔼[X k] + 𝔼[∑ i in s, X i]) ^ 2 : + begin + rw [integral_add', integral_add', integral_add'], + { exact mem_ℒp.integrable ennreal.one_le_two (hs _ (mem_insert_self _ _)) }, + { apply integrable_finset_sum' _ (λ i hi, _), + exact mem_ℒp.integrable ennreal.one_le_two (hs _ (mem_insert_of_mem hi)) }, + { exact mem_ℒp.integrable_sq (hs _ (mem_insert_self _ _)) }, + { apply mem_ℒp.integrable_sq, + exact mem_ℒp_finset_sum' _ (λ i hi, (hs _ (mem_insert_of_mem hi))) }, + { apply integrable.add, + { exact mem_ℒp.integrable_sq (hs _ (mem_insert_self _ _)) }, + { apply mem_ℒp.integrable_sq, + exact mem_ℒp_finset_sum' _ (λ i hi, (hs _ (mem_insert_of_mem hi))) } }, + { rw mul_assoc, + apply integrable.const_mul _ 2, + simp only [mul_sum, sum_apply, pi.mul_apply], + apply integrable_finset_sum _ (λ i hi, _), + apply indep_fun.integrable_mul _ + (mem_ℒp.integrable ennreal.one_le_two (hs _ (mem_insert_self _ _))) + (mem_ℒp.integrable ennreal.one_le_two (hs _ (mem_insert_of_mem hi))), + apply h (mem_insert_self _ _) (mem_insert_of_mem hi), + exact (λ hki, ks (hki.symm ▸ hi)) } + end + ... = Var[X k] + Var[∑ i in s, X i] + + (𝔼[2 * X k * ∑ i in s, X i] - 2 * 𝔼[X k] * 𝔼[∑ i in s, X i]) : + begin + rw [variance_def' (hs _ (mem_insert_self _ _)), + variance_def' (mem_ℒp_finset_sum' _ (λ i hi, (hs _ (mem_insert_of_mem hi))))], + ring, + end + ... = Var[X k] + Var[∑ i in s, X i] : + begin + simp only [mul_assoc, integral_mul_left, pi.mul_apply, pi.bit0_apply, pi.one_apply, sum_apply, + add_right_eq_self, mul_sum], + rw integral_finset_sum s (λ i hi, _), swap, + { apply integrable.const_mul _ 2, + apply indep_fun.integrable_mul _ + (mem_ℒp.integrable ennreal.one_le_two (hs _ (mem_insert_self _ _))) + (mem_ℒp.integrable ennreal.one_le_two (hs _ (mem_insert_of_mem hi))), + apply h (mem_insert_self _ _) (mem_insert_of_mem hi), + exact (λ hki, ks (hki.symm ▸ hi)) }, + rw [integral_finset_sum s + (λ i hi, (mem_ℒp.integrable ennreal.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], + { apply h (mem_insert_self _ _) (mem_insert_of_mem hi), + exact (λ hki, ks (hki.symm ▸ hi)) }, + { exact mem_ℒp.integrable ennreal.one_le_two (hs _ (mem_insert_self _ _)) }, + { exact mem_ℒp.integrable ennreal.one_le_two (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)) + (h.mono (by simp only [coe_insert, set.subset_insert])) +end + +end probability_theory From 5166765790949718234828bf0f974afbbd9feb49 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 7 May 2022 15:50:09 +0000 Subject: [PATCH 12/37] chore(order/filter/pointwise): Better definitional unfolding (#13941) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Tweak pointwise operation definitions to make them easier to work with: * `1` is now `pure 1` instead of `principal 1`. This changes defeq. * Binary operations unfold to the set operation instead exposing a bare `set.image2` (`obtain ⟨t₁, t₂, h₁, h₂, h⟩ : s ∈ f * g` now gives `h : t₁ * t₂ ⊆ s` instead of `h : set.image2 (*) t₁ t₂ ⊆ s`. This does not change defeq. --- src/order/filter/pointwise.lean | 49 +++++++++++++++++++-------------- 1 file changed, 29 insertions(+), 20 deletions(-) diff --git a/src/order/filter/pointwise.lean b/src/order/filter/pointwise.lean index b8517656fcbc0..71da67d815b7e 100644 --- a/src/order/filter/pointwise.lean +++ b/src/order/filter/pointwise.lean @@ -17,8 +17,8 @@ distribute over pointwise operations. For example, ## Main declarations -* `0` (`filter.has_zero`): Principal filter at `0 : α`. -* `1` (`filter.has_one`): Principal filter at `1 : α`. +* `0` (`filter.has_zero`): Pure filter at `0 : α`, or alternatively principal filter at `0 : set α`. +* `1` (`filter.has_one`): Pure filter at `1 : α`, or alternatively principal filter at `1 : set α`. * `f + g` (`filter.has_add`): Addition, filter generated by all `s + t` where `s ∈ f` and `t ∈ g`. * `f * g` (`filter.has_mul`): Multiplication, filter generated by all `s * t` where `s ∈ f` and `t ∈ g`. @@ -63,30 +63,27 @@ variables [has_one α] {f : filter α} {s : set α} /-- `1 : filter α` is defined as the filter of sets containing `1 : α` in locale `pointwise`. -/ @[to_additive "`0 : filter α` is defined as the filter of sets containing `0 : α` in locale `pointwise`."] -protected def has_one : has_one (filter α) := ⟨principal 1⟩ +protected def has_one : has_one (filter α) := ⟨pure 1⟩ localized "attribute [instance] filter.has_one filter.has_zero" in pointwise -@[simp, to_additive] lemma mem_one : s ∈ (1 : filter α) ↔ (1 : α) ∈ s := one_subset - -@[to_additive] lemma one_mem_one : (1 : set α) ∈ (1 : filter α) := mem_principal_self _ - -@[simp, to_additive] lemma principal_one : 𝓟 1 = (1 : filter α) := rfl -@[simp, to_additive] lemma le_one_iff : f ≤ 1 ↔ (1 : set α) ∈ f := le_principal_iff -@[simp, to_additive] lemma pure_one : pure 1 = (1 : filter α) := (principal_singleton _).symm -@[simp, to_additive] lemma eventually_one {p : α → Prop} : (∀ᶠ x in 1, p x) ↔ p 1 := -by rw [←pure_one, eventually_pure] +@[simp, to_additive] lemma mem_one : s ∈ (1 : filter α) ↔ (1 : α) ∈ s := mem_pure +@[to_additive] lemma one_mem_one : (1 : set α) ∈ (1 : filter α) := mem_pure.2 one_mem_one +@[simp, to_additive] lemma pure_one : pure 1 = (1 : filter α) := rfl +@[simp, to_additive] lemma principal_one : 𝓟 1 = (1 : filter α) := principal_singleton _ +@[to_additive] lemma one_ne_bot : (1 : filter α).ne_bot := filter.pure_ne_bot +@[simp, to_additive] protected lemma map_one' (f : α → β) : (1 : filter α).map f = pure (f 1) := rfl +@[simp, to_additive] lemma le_one_iff : f ≤ 1 ↔ (1 : set α) ∈ f := le_pure_iff +@[simp, to_additive] lemma eventually_one {p : α → Prop} : (∀ᶠ x in 1, p x) ↔ p 1 := eventually_pure @[simp, to_additive] lemma tendsto_one {a : filter β} {f : β → α} : tendsto f a 1 ↔ ∀ᶠ x in a, f x = 1 := -by rw [←pure_one, tendsto_pure] +tendsto_pure variables [has_one β] @[simp, to_additive] protected lemma map_one [one_hom_class F α β] (φ : F) : map φ 1 = 1 := -le_antisymm - (le_principal_iff.2 $ mem_map_iff_exists_image.2 ⟨1, one_mem_one, λ x, by simp [map_one φ]⟩) - (le_map $ λ s hs, mem_one.2 ⟨1, mem_one.1 hs, map_one φ⟩) +by rw [filter.map_one', map_one, pure_one] end one @@ -97,7 +94,10 @@ variables [has_mul α] [has_mul β] {f f₁ f₂ g g₁ g₂ h : filter α} {s t /-- The filter `f * g` is generated by `{s * t | s ∈ f, t ∈ g}` in locale `pointwise`. -/ @[to_additive "The filter `f + g` is generated by `{s + t | s ∈ f, t ∈ g}` in locale `pointwise`."] -protected def has_mul : has_mul (filter α) := ⟨map₂ (*)⟩ +protected def has_mul : has_mul (filter α) := +/- This is defeq to `map₂ (*) f g`, but the hypothesis unfolds to `t₁ * t₂ ⊆ s` rather than all the +way to `set.image2 (*) t₁ t₂ ⊆ s`. -/ +⟨λ f g, { sets := {s | ∃ t₁ t₂, t₁ ∈ f ∧ t₂ ∈ g ∧ t₁ * t₂ ⊆ s}, ..map₂ (*) f g }⟩ localized "attribute [instance] filter.has_mul filter.has_add" in pointwise @@ -246,7 +246,10 @@ variables [has_div α] {f f₁ f₂ g g₁ g₂ h : filter α} {s t : set α} {a /-- The filter `f / g` is generated by `{s / t | s ∈ f, t ∈ g}` in locale `pointwise`. -/ @[to_additive "The filter `f - g` is generated by `{s - t | s ∈ f, t ∈ g}` in locale `pointwise`."] -protected def has_div : has_div (filter α) := ⟨map₂ (/)⟩ +protected def has_div : has_div (filter α) := +/- This is defeq to `map₂ (/) f g`, but the hypothesis unfolds to `t₁ / t₂ ⊆ s` rather than all the +way to `set.image2 (/) t₁ t₂ ⊆ s`. -/ +⟨λ f g, { sets := {s | ∃ t₁ t₂, t₁ ∈ f ∧ t₂ ∈ g ∧ t₁ / t₂ ⊆ s}, ..map₂ (/) f g }⟩ localized "attribute [instance] filter.has_div filter.has_sub" in pointwise @@ -318,7 +321,10 @@ section smul variables [has_scalar α β] {f f₁ f₂ : filter α} {g g₁ g₂ h : filter β} {s : set α} {t : set β} {a : α} {b : β} -@[to_additive filter.has_vadd] instance : has_scalar (filter α) (filter β) := ⟨map₂ (•)⟩ +@[to_additive filter.has_vadd] instance : has_scalar (filter α) (filter β) := +/- This is defeq to `map₂ (•) f g`, but the hypothesis unfolds to `t₁ • t₂ ⊆ s` rather than all the +way to `set.image2 (•) t₁ t₂ ⊆ s`. -/ +⟨λ f g, { sets := {s | ∃ t₁ t₂, t₁ ∈ f ∧ t₂ ∈ g ∧ t₁ • t₂ ⊆ s}, ..map₂ (•) f g }⟩ @[simp, to_additive] lemma map₂_smul : map₂ (•) f g = f • g := rfl @[to_additive] lemma mem_smul : t ∈ f • g ↔ ∃ t₁ t₂, t₁ ∈ f ∧ t₂ ∈ g ∧ t₁ • t₂ ⊆ t := iff.rfl @@ -357,7 +363,10 @@ variables [has_vsub α β] {f f₁ f₂ g g₁ g₂ : filter β} {h : filter α} include α /-- The filter `f -ᵥ g` is generated by `{s -ᵥ t | s ∈ f, t ∈ g}` in locale `pointwise`. -/ -protected def has_vsub : has_vsub (filter α) (filter β) := ⟨map₂ (-ᵥ)⟩ +protected def has_vsub : has_vsub (filter α) (filter β) := +/- This is defeq to `map₂ (-ᵥ) f g`, but the hypothesis unfolds to `t₁ -ᵥ t₂ ⊆ s` rather than all +the way to `set.image2 (-ᵥ) t₁ t₂ ⊆ s`. -/ +⟨λ f g, { sets := {s | ∃ t₁ t₂, t₁ ∈ f ∧ t₂ ∈ g ∧ t₁ -ᵥ t₂ ⊆ s}, ..map₂ (-ᵥ) f g }⟩ localized "attribute [instance] filter.has_vsub" in pointwise From 0e494afeeb1945dbc6fec3bdf933e14892f4a018 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 7 May 2022 15:50:10 +0000 Subject: [PATCH 13/37] chore(order/*): Less `order_dual` abuse (#14008) Sanitize uses of `order_dual` by inserting the required `of_dual` and `to_dual` instead of type-ascripting. Also remove some uses which were not necessary. Those dated from the time where we did not have antitone functions. --- .../100-theorems-list/82_cubing_a_cube.lean | 11 +++-- src/algebra/order/field.lean | 8 ++-- src/analysis/convex/quasiconvex.lean | 16 ++------ src/data/finset/lattice.lean | 2 + .../liouville/liouville_constant.lean | 2 +- src/order/compare.lean | 21 +++++----- src/order/galois_connection.lean | 41 +++++++++++-------- src/order/monotone.lean | 2 +- src/order/ord_continuous.lean | 10 ++--- 9 files changed, 56 insertions(+), 57 deletions(-) diff --git a/archive/100-theorems-list/82_cubing_a_cube.lean b/archive/100-theorems-list/82_cubing_a_cube.lean index a96eeefa80ea1..1f94ec93cee2c 100644 --- a/archive/100-theorems-list/82_cubing_a_cube.lean +++ b/archive/100-theorems-list/82_cubing_a_cube.lean @@ -496,13 +496,12 @@ noncomputable def sequence_of_cubes : ℕ → { i : ι // valley cs ((cs i).shif | 0 := let v := valley_unit_cube h in ⟨mi h v, valley_mi⟩ | (k+1) := let v := (sequence_of_cubes k).2 in ⟨mi h v, valley_mi⟩ -def decreasing_sequence (k : ℕ) : order_dual ℝ := -(cs (sequence_of_cubes h k).1).w +def decreasing_sequence (k : ℕ) : ℝ := (cs (sequence_of_cubes h k).1).w -lemma strict_mono_sequence_of_cubes : strict_mono $ decreasing_sequence h := -strict_mono_nat_of_lt_succ $ +lemma strict_anti_sequence_of_cubes : strict_anti $ decreasing_sequence h := +strict_anti_nat_of_succ_lt $ λ k, begin - intro k, let v := (sequence_of_cubes h k).2, dsimp only [decreasing_sequence, sequence_of_cubes], + let v := (sequence_of_cubes h k).2, dsimp only [decreasing_sequence, sequence_of_cubes], apply w_lt_w h v (mi_mem_bcubes : mi h v ∈ _), end @@ -512,7 +511,7 @@ theorem not_correct : ¬correct cs := begin intro h, apply (lt_omega_of_fintype ι).not_le, rw [omega, lift_id], fapply mk_le_of_injective, exact λ n, (sequence_of_cubes h n).1, - intros n m hnm, apply strict_mono.injective (strict_mono_sequence_of_cubes h), + intros n m hnm, apply (strict_anti_sequence_of_cubes h).injective, dsimp only [decreasing_sequence], rw hnm end diff --git a/src/algebra/order/field.lean b/src/algebra/order/field.lean index a2c1521c88c89..c6df55bc62ed6 100644 --- a/src/algebra/order/field.lean +++ b/src/algebra/order/field.lean @@ -706,11 +706,11 @@ eq.symm $ monotone.map_max (λ x y, div_le_div_of_le hc) lemma min_div_div_right_of_nonpos {c : α} (hc : c ≤ 0) (a b : α) : min (a / c) (b / c) = (max a b) / c := -eq.symm $ @monotone.map_max α αᵒᵈ _ _ _ _ _ (λ x y, div_le_div_of_nonpos_of_le hc) +eq.symm $ antitone.map_max $ λ x y, div_le_div_of_nonpos_of_le hc lemma max_div_div_right_of_nonpos {c : α} (hc : c ≤ 0) (a b : α) : max (a / c) (b / c) = (min a b) / c := -eq.symm $ @monotone.map_min α αᵒᵈ _ _ _ _ _ (λ x y, div_le_div_of_nonpos_of_le hc) +eq.symm $ antitone.map_min $ λ x y, div_le_div_of_nonpos_of_le hc lemma abs_div (a b : α) : |a / b| = |a| / |b| := (abs_hom : α →*₀ α).map_div a b @@ -733,10 +733,10 @@ lemma one_div_pow_lt_one_div_pow_of_lt (a1 : 1 < a) {m n : ℕ} (mn : m < n) : by refine (one_div_lt_one_div _ _).mpr (pow_lt_pow a1 mn); exact pow_pos (trans zero_lt_one a1) _ -lemma one_div_pow_mono (a1 : 1 ≤ a) : monotone (λ n : ℕ, order_dual.to_dual 1 / a ^ n) := +lemma one_div_pow_anti (a1 : 1 ≤ a) : antitone (λ n : ℕ, 1 / a ^ n) := λ m n, one_div_pow_le_one_div_pow_of_le a1 -lemma one_div_pow_strict_mono (a1 : 1 < a) : strict_mono (λ n : ℕ, order_dual.to_dual 1 / a ^ n) := +lemma one_div_pow_strict_anti (a1 : 1 < a) : strict_anti (λ n : ℕ, 1 / a ^ n) := λ m n, one_div_pow_lt_one_div_pow_of_lt a1 /-! ### Results about `is_lub` and `is_glb` -/ diff --git a/src/analysis/convex/quasiconvex.lean b/src/analysis/convex/quasiconvex.lean index 6fd19188f29d9..a89dceac86ea5 100644 --- a/src/analysis/convex/quasiconvex.lean +++ b/src/analysis/convex/quasiconvex.lean @@ -31,7 +31,7 @@ not hard but quite a pain to go about as there are many cases to consider. * https://en.wikipedia.org/wiki/Quasiconvex_function -/ -open function set +open function order_dual set variables {𝕜 E F β : Type*} @@ -62,17 +62,9 @@ quasiconvex_on 𝕜 s f ∧ quasiconcave_on 𝕜 s f variables {𝕜 s f} -lemma quasiconvex_on.dual (hf : quasiconvex_on 𝕜 s f) : - @quasiconcave_on 𝕜 E (order_dual β) _ _ _ _ s f := -hf - -lemma quasiconcave_on.dual (hf : quasiconcave_on 𝕜 s f) : - @quasiconvex_on 𝕜 E (order_dual β) _ _ _ _ s f := -hf - -lemma quasilinear_on.dual (hf : quasilinear_on 𝕜 s f) : - @quasilinear_on 𝕜 E (order_dual β) _ _ _ _ s f := -⟨hf.2, hf.1⟩ +lemma quasiconvex_on.dual : quasiconvex_on 𝕜 s f → quasiconcave_on 𝕜 s (to_dual ∘ f) := id +lemma quasiconcave_on.dual : quasiconcave_on 𝕜 s f → quasiconvex_on 𝕜 s (to_dual ∘ f) := id +lemma quasilinear_on.dual : quasilinear_on 𝕜 s f → quasilinear_on 𝕜 s (to_dual ∘ f) := and.swap lemma convex.quasiconvex_on_of_convex_le (hs : convex 𝕜 s) (h : ∀ r, convex 𝕜 {x | f x ≤ r}) : quasiconvex_on 𝕜 s f := diff --git a/src/data/finset/lattice.lean b/src/data/finset/lattice.lean index a594502965903..7e275fba769c7 100644 --- a/src/data/finset/lattice.lean +++ b/src/data/finset/lattice.lean @@ -607,6 +607,8 @@ lemma inf'_le (h : b ∈ s) : s.inf' ⟨b, h⟩ f ≤ f b := @le_sup' αᵒᵈ _ @[simp] lemma inf'_const (a : α) : s.inf' H (λ b, a) = a := @sup'_const αᵒᵈ _ _ _ _ _ +@[simp] lemma le_inf'_iff : a ≤ s.inf' H f ↔ ∀ b ∈ s, a ≤ f b := @sup'_le_iff αᵒᵈ _ _ _ H f _ + lemma inf'_bUnion [decidable_eq β] {s : finset γ} (Hs : s.nonempty) {t : γ → finset β} (Ht : ∀ b, (t b).nonempty) : (s.bUnion t).inf' (Hs.bUnion (λ b _, Ht b)) f = s.inf' Hs (λ b, (t b).inf' (Ht b) f) := diff --git a/src/number_theory/liouville/liouville_constant.lean b/src/number_theory/liouville/liouville_constant.lean index f641942477311..981d728046b42 100644 --- a/src/number_theory/liouville/liouville_constant.lean +++ b/src/number_theory/liouville/liouville_constant.lean @@ -106,7 +106,7 @@ calc (∑' i, 1 / m ^ (i + (n + 1))!) (λ b, one_div_pow_le_one_div_pow_of_le m1.le (b.add_factorial_succ_le_factorial_add_succ n)) -- 3. the term with index `i = 2` of the first series is strictly smaller than -- the corresponding term of the second series - (one_div_pow_strict_mono m1 (n.add_factorial_succ_lt_factorial_add_succ rfl.le)) + (one_div_pow_strict_anti m1 (n.add_factorial_succ_lt_factorial_add_succ rfl.le)) -- 4. the second series is summable, since its terms grow quickly (summable_one_div_pow_of_le m1 (λ j, nat.le.intro rfl)) ... = ∑' i, (1 / m) ^ i * (1 / m ^ (n + 1)!) : diff --git a/src/order/compare.lean b/src/order/compare.lean index a9cefd8331663..f8d35efda182f 100644 --- a/src/order/compare.lean +++ b/src/order/compare.lean @@ -3,7 +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 order.basic +import order.order_dual /-! # Comparison @@ -126,22 +126,23 @@ by cases o₁; cases o₂; exact dec_trivial end ordering -lemma order_dual.dual_compares [has_lt α] {a b : α} {o : ordering} : - @ordering.compares αᵒᵈ _ o a b ↔ @ordering.compares α _ o b a := +open ordering order_dual + +@[simp] lemma to_dual_compares_to_dual [has_lt α] {a b : α} {o : ordering} : + compares o (to_dual a) (to_dual b) ↔ compares o b a := +by { cases o, exacts [iff.rfl, eq_comm, iff.rfl] } + +@[simp] lemma of_dual_compares_of_dual [has_lt α] {a b : αᵒᵈ} {o : ordering} : + compares o (of_dual a) (of_dual b) ↔ compares o b a := by { cases o, exacts [iff.rfl, eq_comm, iff.rfl] } lemma cmp_compares [linear_order α] (a b : α) : (cmp a b).compares a b := -begin - unfold cmp cmp_using, - by_cases a < b; simp [h], - by_cases h₂ : b < a; simp [h₂, gt], - exact (decidable.lt_or_eq_of_le (le_of_not_gt h₂)).resolve_left h -end +by obtain h | h | h := lt_trichotomy a b; simp [cmp, cmp_using, h, h.not_lt] lemma cmp_swap [preorder α] [@decidable_rel α (<)] (a b : α) : (cmp a b).swap = cmp b a := begin unfold cmp cmp_using, - by_cases a < b; by_cases h₂ : b < a; simp [h, h₂, gt, ordering.swap], + by_cases a < b; by_cases h₂ : b < a; simp [h, h₂, ordering.swap], exact lt_asymm h h₂ end diff --git a/src/order/galois_connection.lean b/src/order/galois_connection.lean index aa514caa681c0..94180f9623d94 100644 --- a/src/order/galois_connection.lean +++ b/src/order/galois_connection.lean @@ -41,7 +41,8 @@ This means the infimum of subgroups will be defined to be the intersection of se with a proof that intersection of subgroups is a subgroup, rather than the closure of the intersection. -/ -open function order set + +open function order_dual set universes u v w x variables {α : Type u} {β : Type v} {γ : Type w} {ι : Sort x} {κ : ι → Sort*} {a a₁ a₂ : α} @@ -359,7 +360,7 @@ def galois_insertion.monotone_intro {α β : Type*} [preorder α] [preorder β] /-- Makes a Galois insertion from an order-preserving bijection. -/ protected def order_iso.to_galois_insertion [preorder α] [preorder β] (oi : α ≃o β) : -@galois_insertion α β _ _ (oi) (oi.symm) := + galois_insertion oi oi.symm := { choice := λ b h, oi b, gc := oi.to_galois_connection, le_l_u := λ g, le_of_eq (oi.right_inv g).symm, @@ -535,35 +536,39 @@ end galois_insertion choice function, to give better definitional equalities when lifting order structures. Dual to `galois_insertion` -/ @[nolint has_inhabited_instance] -structure galois_coinsertion {α β : Type*} [preorder α] [preorder β] (l : α → β) (u : β → α) := +structure galois_coinsertion [preorder α] [preorder β] (l : α → β) (u : β → α) := (choice : Πx : β, x ≤ l (u x) → α) (gc : galois_connection l u) (u_l_le : ∀ x, u (l x) ≤ x) (choice_eq : ∀ a h, choice a h = u a) -/-- Make a `galois_insertion u l` in the `order_dual`, from a `galois_coinsertion l u` -/ -def galois_coinsertion.dual {α β : Type*} [preorder α] [preorder β] {l : α → β} {u : β → α} : - galois_coinsertion l u → @galois_insertion (order_dual β) (order_dual α) _ _ u l := +/-- Make a `galois_insertion` between `αᵒᵈ` and `βᵒᵈ` from a `galois_coinsertion` between `α` and +`β`. -/ +def galois_coinsertion.dual [preorder α] [preorder β] {l : α → β} {u : β → α} : + galois_coinsertion l u → galois_insertion (to_dual ∘ u ∘ of_dual) (to_dual ∘ l ∘ of_dual) := λ x, ⟨x.1, x.2.dual, x.3, x.4⟩ -/-- Make a `galois_coinsertion u l` in the `order_dual`, from a `galois_insertion l u` -/ -def galois_insertion.dual {α β : Type*} [preorder α] [preorder β] {l : α → β} {u : β → α} : - galois_insertion l u → @galois_coinsertion (order_dual β) (order_dual α) _ _ u l := +/-- Make a `galois_coinsertion` between `αᵒᵈ` and `βᵒᵈ` from a `galois_insertion` between `α` and +`β`. -/ +def galois_insertion.dual [preorder α] [preorder β] {l : α → β} {u : β → α} : + galois_insertion l u → galois_coinsertion (to_dual ∘ u ∘ of_dual) (to_dual ∘ l ∘ of_dual) := λ x, ⟨x.1, x.2.dual, x.3, x.4⟩ -/-- Make a `galois_coinsertion l u` from a `galois_insertion l u` in the `order_dual` -/ -def galois_coinsertion.of_dual {α β : Type*} [preorder α] [preorder β] {l : α → β} {u : β → α} : - @galois_insertion (order_dual β) (order_dual α) _ _ u l → galois_coinsertion l u := +/-- Make a `galois_insertion` between `α` and `β` from a `galois_coinsertion` between `αᵒᵈ` and +`βᵒᵈ`. -/ +def galois_coinsertion.of_dual [preorder α] [preorder β] {l : αᵒᵈ → βᵒᵈ} {u : βᵒᵈ → αᵒᵈ} : + galois_coinsertion l u → galois_insertion (of_dual ∘ u ∘ to_dual) (of_dual ∘ l ∘ to_dual) := λ x, ⟨x.1, x.2.dual, x.3, x.4⟩ -/-- Make a `galois_insertion l u` from a `galois_coinsertion l u` in the `order_dual` -/ -def galois_insertion.of_dual {α β : Type*} [preorder α] [preorder β] {l : α → β} {u : β → α} : - @galois_coinsertion (order_dual β) (order_dual α) _ _ u l → galois_insertion l u := +/-- Make a `galois_coinsertion` between `α` and `β` from a `galois_insertion` between `αᵒᵈ` and +`βᵒᵈ`. -/ +def galois_insertion.of_dual [preorder α] [preorder β] {l : αᵒᵈ → βᵒᵈ} {u : βᵒᵈ → αᵒᵈ} : + galois_insertion l u → galois_coinsertion (of_dual ∘ u ∘ to_dual) (of_dual ∘ l ∘ to_dual) := λ x, ⟨x.1, x.2.dual, x.3, x.4⟩ /-- Makes a Galois coinsertion from an order-preserving bijection. -/ -protected def rel_iso.to_galois_coinsertion [preorder α] [preorder β] (oi : α ≃o β) : -@galois_coinsertion α β _ _ (oi) (oi.symm) := +protected def order_iso.to_galois_coinsertion [preorder α] [preorder β] (oi : α ≃o β) : + galois_coinsertion oi oi.symm := { choice := λ b h, oi.symm b, gc := oi.to_galois_connection, u_l_le := λ g, le_of_eq (oi.left_inv g), @@ -574,7 +579,7 @@ def galois_coinsertion.monotone_intro [preorder α] [preorder β] {l : α → β (hu : monotone u) (hl : monotone l) (hlu : ∀ b, l (u b) ≤ b) (hul : ∀ a, u (l a) = a) : galois_coinsertion l u := -galois_coinsertion.of_dual (galois_insertion.monotone_intro hl.dual hu.dual hlu hul) +(galois_insertion.monotone_intro hl.dual hu.dual hlu hul).of_dual /-- Make a `galois_coinsertion l u` from a `galois_connection l u` such that `∀ b, b ≤ l (u b)` -/ def galois_connection.to_galois_coinsertion {α β : Type*} [preorder α] [preorder β] diff --git a/src/order/monotone.lean b/src/order/monotone.lean index 57f24e1e1f547..e28880513f4b7 100644 --- a/src/order/monotone.lean +++ b/src/order/monotone.lean @@ -557,7 +557,7 @@ protected theorem strict_mono_on.compares (hf : strict_mono_on f s) {a b : α} ( protected theorem strict_anti_on.compares (hf : strict_anti_on f s) {a b : α} (ha : a ∈ s) (hb : b ∈ s) {o : ordering} : o.compares (f a) (f b) ↔ o.compares b a := -order_dual.dual_compares.trans $ hf.dual_right.compares hb ha +to_dual_compares_to_dual.trans $ hf.dual_right.compares hb ha protected theorem strict_mono.compares (hf : strict_mono f) {a b : α} {o : ordering} : o.compares (f a) (f b) ↔ o.compares a b := diff --git a/src/order/ord_continuous.lean b/src/order/ord_continuous.lean index f51b9b3256f95..f89d5c73a0fc7 100644 --- a/src/order/ord_continuous.lean +++ b/src/order/ord_continuous.lean @@ -23,7 +23,7 @@ universes u v w x variables {α : Type u} {β : Type v} {γ : Type w} {ι : Sort x} -open set function +open function order_dual set /-! ### Definitions @@ -51,8 +51,8 @@ protected lemma id : left_ord_continuous (id : α → α) := λ s x h, by simpa variable {α} -protected lemma order_dual (hf : left_ord_continuous f) : - @right_ord_continuous (order_dual α) (order_dual β) _ _ f := hf +protected lemma order_dual : left_ord_continuous f → right_ord_continuous (to_dual ∘ f ∘ of_dual) := +id lemma map_is_greatest (hf : left_ord_continuous f) {s : set α} {x : α} (h : is_greatest s x): is_greatest (f '' s) (f x) := @@ -148,8 +148,8 @@ protected lemma id : right_ord_continuous (id : α → α) := λ s x h, by simpa variable {α} -protected lemma order_dual (hf : right_ord_continuous f) : - @left_ord_continuous (order_dual α) (order_dual β) _ _ f := hf +protected lemma order_dual : right_ord_continuous f → left_ord_continuous (to_dual ∘ f ∘ of_dual) := +id lemma map_is_least (hf : right_ord_continuous f) {s : set α} {x : α} (h : is_least s x): is_least (f '' s) (f x) := From e0dd30031013653bccdac0a4fa58e413453e4678 Mon Sep 17 00:00:00 2001 From: damiano Date: Sat, 7 May 2022 20:15:08 +0000 Subject: [PATCH 14/37] =?UTF-8?q?feat(algebra/{invertible=20+=20group=5Fpo?= =?UTF-8?q?wer/lemmas}):=20taking=20`inv=5Fof`=20(=E2=85=9F=5F)=20is=20inj?= =?UTF-8?q?ective=20(#14011)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Besides the lemma stated in the description, I also made explicit an argument that was implicit in a different lemma and swapped the arguments of `invertible_unique` in order to get the typeclass assumptions before some non-typeclass assumptions. [Zulip](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/inv_of_inj) Co-authored-by: Floris van Doorn <[fpvdoorn@gmail.com](mailto:fpvdoorn@gmail.com)> --- src/algebra/group_power/lemmas.lean | 2 +- src/algebra/invertible.lean | 10 +++++++--- 2 files changed, 8 insertions(+), 4 deletions(-) diff --git a/src/algebra/group_power/lemmas.lean b/src/algebra/group_power/lemmas.lean index 2410fcf2a94bc..56a4139af2701 100644 --- a/src/algebra/group_power/lemmas.lean +++ b/src/algebra/group_power/lemmas.lean @@ -41,7 +41,7 @@ instance invertible_pow (m : M) [invertible m] (n : ℕ) : invertible (m ^ n) := lemma inv_of_pow (m : M) [invertible m] (n : ℕ) [invertible (m ^ n)] : ⅟(m ^ n) = ⅟m ^ n := -@invertible_unique M _ (m ^ n) (m ^ n) rfl ‹_› (invertible_pow m n) +@invertible_unique M _ (m ^ n) (m ^ n) _ (invertible_pow m n) rfl lemma is_unit.pow {m : M} (n : ℕ) : is_unit m → is_unit (m ^ n) := λ ⟨u, hu⟩, ⟨u ^ n, by simp *⟩ diff --git a/src/algebra/invertible.lean b/src/algebra/invertible.lean index 19ebbb9f01ec6..f1ab5c27eaf50 100644 --- a/src/algebra/invertible.lean +++ b/src/algebra/invertible.lean @@ -97,8 +97,8 @@ left_inv_eq_right_inv (inv_of_mul_self _) hac lemma inv_of_eq_left_inv [monoid α] {a b : α} [invertible a] (hac : b * a = 1) : ⅟a = b := (left_inv_eq_right_inv hac (mul_inv_of_self _)).symm -lemma invertible_unique {α : Type u} [monoid α] (a b : α) (h : a = b) - [invertible a] [invertible b] : +lemma invertible_unique {α : Type u} [monoid α] (a b : α) [invertible a] [invertible b] + (h : a = b) : ⅟a = ⅟b := by { apply inv_of_eq_right_inv, rw [h, mul_inv_of_self], } @@ -177,10 +177,14 @@ by simp only [←two_mul, mul_inv_of_self] instance invertible_inv_of [has_one α] [has_mul α] {a : α} [invertible a] : invertible (⅟a) := ⟨ a, mul_inv_of_self a, inv_of_mul_self a ⟩ -@[simp] lemma inv_of_inv_of [monoid α] {a : α} [invertible a] [invertible (⅟a)] : +@[simp] lemma inv_of_inv_of [monoid α] (a : α) [invertible a] [invertible (⅟a)] : ⅟(⅟a) = a := inv_of_eq_right_inv (inv_of_mul_self _) +@[simp] lemma inv_of_inj [monoid α] {a b : α} [invertible a] [invertible b] : + ⅟ a = ⅟ b ↔ a = b := +⟨invertible_unique _ _, invertible_unique _ _⟩ + /-- `⅟b * ⅟a` is the inverse of `a * b` -/ def invertible_mul [monoid α] (a b : α) [invertible a] [invertible b] : invertible (a * b) := ⟨ ⅟b * ⅟a, by simp [←mul_assoc], by simp [←mul_assoc] ⟩ From e1902256fb5248e57ec9fc08b9ac12a98ee9e426 Mon Sep 17 00:00:00 2001 From: Rob Lewis Date: Sat, 7 May 2022 22:21:47 +0000 Subject: [PATCH 15/37] feat(data/rat/meta_defs, meta/expr): rat.to_pexpr and int.to_pexpr (#14002) Co-authored-by: Rob Lewis --- src/data/rat/meta_defs.lean | 11 +++++++++++ src/meta/expr.lean | 10 ++++++++++ 2 files changed, 21 insertions(+) diff --git a/src/data/rat/meta_defs.lean b/src/data/rat/meta_defs.lean index 6cae57de5c0ba..abff0313e2258 100644 --- a/src/data/rat/meta_defs.lean +++ b/src/data/rat/meta_defs.lean @@ -53,6 +53,17 @@ local attribute [semireducible] reflected meta instance : has_reflect ℚ := rat.reflect end +/-- +`rat.to_pexpr q` creates a `pexpr` that will evaluate to `q`. +The `pexpr` does not hold any typing information: +`to_expr ``((%%(rat.to_pexpr (3/4)) : K))` will create a native `K` numeral `(3/4 : K)`. +-/ +meta def rat.to_pexpr (q : ℚ) : pexpr := +let n := q.num, + d := q.denom in +if d = 1 then n.to_pexpr +else ``(%%n.to_pexpr / %%d.to_pexpr) + /-- Evaluates an expression as a rational number, if that expression represents a numeral or the quotient of two numerals. -/ protected meta def expr.to_nonneg_rat : expr → option ℚ diff --git a/src/meta/expr.lean b/src/meta/expr.lean index 8ba3871115634..9d2c045dca993 100644 --- a/src/meta/expr.lean +++ b/src/meta/expr.lean @@ -326,6 +326,16 @@ meta def nat.to_pexpr : ℕ → pexpr | 0 := ``(0) | 1 := ``(1) | n := if n % 2 = 0 then ``(bit0 %%(nat.to_pexpr (n/2))) else ``(bit1 %%(nat.to_pexpr (n/2))) + +/-- +`int.to_pexpr n` creates a `pexpr` that will evaluate to `n`. +The `pexpr` does not hold any typing information: +`to_expr ``((%%(int.to_pexpr (-5)) : ℚ))` will create a native `ℚ` numeral `(-5 : ℚ)`. +-/ +meta def int.to_pexpr : ℤ → pexpr +| (int.of_nat k) := k.to_pexpr +| (int.neg_succ_of_nat k) := ``(-%%((k+1).to_pexpr)) + namespace expr /-- From 3a0eb4b5606be0c4f0d51fdc5c726925b8120b25 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sat, 7 May 2022 22:21:48 +0000 Subject: [PATCH 16/37] chore(logic/relation): Dot notation on `well_founded.trans_gen` (#14016) --- src/logic/relation.lean | 2 +- src/set_theory/game/pgame.lean | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/src/logic/relation.lean b/src/logic/relation.lean index ea07434f1c40b..588181c272806 100644 --- a/src/logic/relation.lean +++ b/src/logic/relation.lean @@ -373,7 +373,7 @@ end end trans_gen -lemma well_founded.trans_gen {α} {r : α → α → Prop} (h : well_founded r) : +lemma _root_.well_founded.trans_gen {α} {r : α → α → Prop} (h : well_founded r) : well_founded (trans_gen r) := ⟨λ a, h.induction a (λ x H, acc.intro x (λ y hy, begin cases hy with _ hyx z _ hyz hzx, diff --git a/src/set_theory/game/pgame.lean b/src/set_theory/game/pgame.lean index 0498c72526ba3..00924a6c460fb 100644 --- a/src/set_theory/game/pgame.lean +++ b/src/set_theory/game/pgame.lean @@ -215,7 +215,7 @@ instance : is_trans _ subsequent := trans_gen.is_trans @[trans] theorem subsequent.trans : ∀ {x y z}, subsequent x y → subsequent y z → subsequent x z := @trans_gen.trans _ _ -theorem wf_subsequent : well_founded subsequent := well_founded.trans_gen wf_is_option +theorem wf_subsequent : well_founded subsequent := wf_is_option.trans_gen instance : has_well_founded pgame := ⟨_, wf_subsequent⟩ From ce0dc83676ba6ce32c3fb32c633f499b89e2fc6f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sun, 8 May 2022 04:50:21 +0000 Subject: [PATCH 17/37] feat(set_theory/ordinal/basic): Supremum indexed over an empty / unique type (#13735) This PR contains the following changes: - The lemmas `sup_unique`, `bsup_one`, `lsub_unique`, `blsub_one`. - `congr` lemmas for `bsup` and `blsub` - Arguments like `o = 0` are removed as the `congr` lemmas now handle this. - `a + 1` is changed to `a.succ` in some lemmas (for better rewriting). --- src/set_theory/cardinal/cofinality.lean | 2 +- src/set_theory/game/birthday.lean | 3 +- src/set_theory/ordinal/arithmetic.lean | 38 +++++++++++++++++++------ 3 files changed, 32 insertions(+), 11 deletions(-) diff --git a/src/set_theory/cardinal/cofinality.lean b/src/set_theory/cardinal/cofinality.lean index ba23e02b7568c..f20ceedf8c4d3 100644 --- a/src/set_theory/cardinal/cofinality.lean +++ b/src/set_theory/cardinal/cofinality.lean @@ -408,7 +408,7 @@ theorem is_fundamental_sequence_id_of_le_cof (h : o ≤ o.cof.ord) : theorem is_fundamental_sequence_zero {f : Π b < (0 : ordinal), ordinal} : is_fundamental_sequence 0 0 f := -⟨by rw [cof_zero, ord_zero], λ i j hi, (ordinal.not_lt_zero i hi).elim, blsub_zero rfl f⟩ +⟨by rw [cof_zero, ord_zero], λ i j hi, (ordinal.not_lt_zero i hi).elim, blsub_zero f⟩ theorem is_fundamental_sequence_succ : is_fundamental_sequence o.succ 1 (λ _ _, o) := begin diff --git a/src/set_theory/game/birthday.lean b/src/set_theory/game/birthday.lean index 673f4f16d2c49..6cef0d34fa933 100644 --- a/src/set_theory/game/birthday.lean +++ b/src/set_theory/game/birthday.lean @@ -104,7 +104,8 @@ by { rw birthday_eq_zero, split; apply_instance } @[simp] theorem birthday_one : birthday 1 = 1 := begin have : (λ i, (move_left 1 i).birthday) = λ i, 0 := funext (λ x, by simp), - rw [birthday_def, @ordinal.lsub_empty (right_moves 1), this, ordinal.lsub_const, zero_add], + rw [birthday_def, @ordinal.lsub_empty (right_moves 1), this, ordinal.lsub_const, + ordinal.succ_zero], exact max_bot_right 1 end diff --git a/src/set_theory/ordinal/arithmetic.lean b/src/set_theory/ordinal/arithmetic.lean index dd84a0e9ccb43..307423455de80 100644 --- a/src/set_theory/ordinal/arithmetic.lean +++ b/src/set_theory/ordinal/arithmetic.lean @@ -1097,14 +1097,17 @@ eq_of_forall_ge_iff $ λ a, by rw [sup_le_iff, comp, H.le_set' (λ_:ι, true) g (let ⟨i⟩ := h in ⟨i, ⟨⟩⟩)]; intros; simp only [sup_le_iff, true_implies_iff]; tauto -theorem sup_empty {ι} [is_empty ι] (f : ι → ordinal) : sup f = 0 := +@[simp] theorem sup_empty {ι} [is_empty ι] (f : ι → ordinal) : sup f = 0 := sup_eq_zero_iff.2 is_empty_elim theorem sup_ord {ι} (f : ι → cardinal) : sup (λ i, (f i).ord) = (cardinal.sup f).ord := eq_of_forall_ge_iff $ λ a, by simp only [sup_le_iff, cardinal.ord_le, cardinal.sup_le_iff] -theorem sup_const {ι} [hι : nonempty ι] (o : ordinal) : sup (λ _ : ι, o) = o := -le_antisymm (sup_le (λ _, le_rfl)) (le_sup _ hι.some) +@[simp] theorem sup_const {ι} [hι : nonempty ι] (o : ordinal) : sup (λ _ : ι, o) = o := +csupr_const + +@[simp] theorem sup_unique {ι} [unique ι] (f : ι → ordinal) : sup f = f default := +supr_unique theorem sup_le_of_range_subset {ι ι'} {f : ι → ordinal} {g : ι' → ordinal} (h : set.range f ⊆ set.range g) : sup.{u (max v w)} f ≤ sup.{v (max u w)} g := @@ -1182,6 +1185,10 @@ by rw [bsup_eq_sup', bsup_eq_sup'] @[simp] theorem bsup_eq_sup {ι} (f : ι → ordinal) : bsup _ (bfamily_of_family f) = sup f := bsup_eq_sup' _ f +@[congr] lemma bsup_congr {o₁ o₂ : ordinal} (f : Π a < o₁, ordinal) (ho : o₁ = o₂) : + bsup o₁ f = bsup o₂ (λ a h, f a (h.trans_eq ho.symm)) := +by subst ho + theorem bsup_le_iff {o f a} : bsup.{u v} o f ≤ a ↔ ∀ i h, f i h ≤ a := sup_le_iff.trans ⟨λ h i hi, by { rw ←family_of_bfamily_enum o f, exact h _ }, λ h i, h _ _⟩ @@ -1222,12 +1229,15 @@ theorem lt_bsup_of_limit {o : ordinal} {f : Π a < o, ordinal} (ho : ∀ a < o, succ a < o) (i h) : f i h < bsup o f := (hf _ _ $ lt_succ_self i).trans_le (le_bsup f i.succ $ ho _ h) -theorem bsup_zero {o : ordinal} (ho : o = 0) (f : Π a < o, ordinal) : bsup o f = 0 := -bsup_eq_zero_iff.2 (λ i hi, by { subst ho, exact (ordinal.not_lt_zero i hi).elim }) +@[simp] theorem bsup_zero (f : Π a < (0 : ordinal), ordinal) : bsup 0 f = 0 := +bsup_eq_zero_iff.2 (λ i hi, (ordinal.not_lt_zero i hi).elim) theorem bsup_const {o : ordinal} (ho : o ≠ 0) (a : ordinal) : bsup o (λ _ _, a) = a := le_antisymm (bsup_le (λ _ _, le_rfl)) (le_bsup _ 0 (ordinal.pos_iff_ne_zero.2 ho)) +@[simp] theorem bsup_one (f : Π a < (1 : ordinal), ordinal) : bsup 1 f = f 0 zero_lt_one := +by simp_rw [←sup_eq_bsup, sup_unique, family_of_bfamily, family_of_bfamily', typein_one_out] + theorem bsup_le_of_brange_subset {o o'} {f : Π a < o, ordinal} {g : Π a < o', ordinal} (h : brange o f ⊆ brange o' g) : bsup.{u (max v w)} o f ≤ bsup.{v (max u w)} o' g := bsup_le $ λ i hi, begin @@ -1302,7 +1312,7 @@ end theorem sup_eq_lsub_iff_lt_sup {ι} (f : ι → ordinal) : sup f = lsub f ↔ ∀ i, f i < sup f := ⟨λ h i, (by { rw h, apply lt_lsub }), λ h, le_antisymm (sup_le_lsub f) (lsub_le h)⟩ -lemma lsub_empty {ι} [h : is_empty ι] (f : ι → ordinal) : lsub f = 0 := +@[simp] lemma lsub_empty {ι} [h : is_empty ι] (f : ι → ordinal) : lsub f = 0 := by { rw [←ordinal.le_zero, lsub_le_iff], exact h.elim } lemma lsub_pos {ι} [h : nonempty ι] (f : ι → ordinal) : 0 < lsub f := @@ -1316,9 +1326,12 @@ begin exact this.false end -theorem lsub_const {ι} [hι : nonempty ι] (o : ordinal) : lsub (λ _ : ι, o) = o + 1 := +@[simp] theorem lsub_const {ι} [hι : nonempty ι] (o : ordinal) : lsub (λ _ : ι, o) = o.succ := sup_const o.succ +@[simp] theorem lsub_unique {ι} [hι : unique ι] (f : ι → ordinal) : lsub f = (f default).succ := +sup_unique _ + theorem lsub_le_of_range_subset {ι ι'} {f : ι → ordinal} {g : ι' → ordinal} (h : set.range f ⊆ set.range g) : lsub.{u (max v w)} f ≤ lsub.{v (max u w)} g := sup_le_of_range_subset (by convert set.image_subset _ h; apply set.range_comp) @@ -1392,6 +1405,10 @@ by rw [blsub_eq_lsub', blsub_eq_lsub'] @[simp] theorem blsub_eq_lsub {ι} (f : ι → ordinal) : blsub _ (bfamily_of_family f) = lsub f := blsub_eq_lsub' _ _ +@[congr] lemma blsub_congr {o₁ o₂ : ordinal} (f : Π a < o₁, ordinal) (ho : o₁ = o₂) : + blsub o₁ f = blsub o₂ (λ a h, f a (h.trans_eq ho.symm)) := +by subst ho + theorem blsub_le_iff {o f a} : blsub o f ≤ a ↔ ∀ i h, f i h < a := by { convert bsup_le_iff, apply propext, simp [succ_le] } @@ -1448,7 +1465,7 @@ end @[simp] theorem blsub_eq_zero_iff {o} {f : Π a < o, ordinal} : blsub o f = 0 ↔ o = 0 := by { rw [←lsub_eq_blsub, lsub_eq_zero_iff], exact out_empty_iff_eq_zero } -lemma blsub_zero {o : ordinal} (ho : o = 0) (f : Π a < o, ordinal) : blsub o f = 0 := +@[simp] lemma blsub_zero (f : Π a < (0 : ordinal), ordinal) : blsub 0 f = 0 := by rwa blsub_eq_zero_iff lemma blsub_pos {o : ordinal} (ho : 0 < o) (f : Π a < o, ordinal) : 0 < blsub o f := @@ -1460,9 +1477,12 @@ eq_of_forall_ge_iff $ λ o, by rw [blsub_le_iff, lsub_le_iff]; exact ⟨λ H b, H _ _, λ H i h, by simpa only [typein_enum] using H (enum r i h)⟩ -theorem blsub_const {o : ordinal} (ho : o ≠ 0) (a : ordinal) : blsub.{u v} o (λ _ _, a) = a + 1 := +theorem blsub_const {o : ordinal} (ho : o ≠ 0) (a : ordinal) : blsub.{u v} o (λ _ _, a) = a.succ := bsup_const.{u v} ho a.succ +@[simp] theorem blsub_one (f : Π a < (1 : ordinal), ordinal) : blsub 1 f = (f 0 zero_lt_one).succ := +bsup_one _ + @[simp] theorem blsub_id : ∀ o, blsub.{u u} o (λ x _, x) = o := lsub_typein From 0c64b3d3c568ab32b344d89b0ee5a799e3e09c4d Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sun, 8 May 2022 06:24:13 +0000 Subject: [PATCH 18/37] feat(algebra/category/Module): biproducts (#13908) Following the same pattern for `AddCommGroup`, create the instance for biproducts in` Module R`, and check they are isomorphic to the usual construction. Co-authored-by: Scott Morrison --- src/algebra/category/Group/biproducts.lean | 43 ++++--- src/algebra/category/Module/biproducts.lean | 129 ++++++++++++++++++++ 2 files changed, 152 insertions(+), 20 deletions(-) create mode 100644 src/algebra/category/Module/biproducts.lean diff --git a/src/algebra/category/Group/biproducts.lean b/src/algebra/category/Group/biproducts.lean index 77988d003a749..40e72a1a6df96 100644 --- a/src/algebra/category/Group/biproducts.lean +++ b/src/algebra/category/Group/biproducts.lean @@ -21,9 +21,20 @@ universe u namespace AddCommGroup +-- As `AddCommGroup` is preadditive, and has all limits, it automatically has biproducts. +instance : has_binary_biproducts AddCommGroup := +has_binary_biproducts.of_has_binary_products + +instance : has_finite_biproducts AddCommGroup := +has_finite_biproducts.of_has_finite_products + +-- We now construct explicit limit data, +-- so we can compare the biproducts to the usual unbundled constructions. + /-- Construct limit data for a binary product in `AddCommGroup`, using `AddCommGroup.of (G × H)`. -/ +@[simps cone_X is_limit_lift] def binary_product_limit_cone (G H : AddCommGroup.{u}) : limits.limit_cone (pair G H) := { cone := { X := AddCommGroup.of (G × H), @@ -36,18 +47,17 @@ def binary_product_limit_cone (G H : AddCommGroup.{u}) : limits.limit_cone (pair ext; [rw ← w walking_pair.left, rw ← w walking_pair.right]; refl, end, } } +@[simp] lemma binary_product_limit_cone_cone_π_app_left (G H : AddCommGroup.{u}) : + (binary_product_limit_cone G H).cone.π.app walking_pair.left = add_monoid_hom.fst G H := rfl -instance has_binary_product (G H : AddCommGroup.{u}) : has_binary_product G H := -has_limit.mk (binary_product_limit_cone G H) - -instance (G H : AddCommGroup.{u}) : has_binary_biproduct G H := -has_binary_biproduct.of_has_binary_product _ _ +@[simp] lemma binary_product_limit_cone_cone_π_app_right (G H : AddCommGroup.{u}) : + (binary_product_limit_cone G H).cone.π.app walking_pair.right = add_monoid_hom.snd G H := rfl /-- We verify that the biproduct in AddCommGroup is isomorphic to the cartesian product of the underlying types: -/ -noncomputable +@[simps] noncomputable def biprod_iso_prod (G H : AddCommGroup.{u}) : (G ⊞ H : AddCommGroup) ≅ AddCommGroup.of (G × H) := is_limit.cone_point_unique_up_to_iso (binary_biproduct.is_limit G H) @@ -75,7 +85,7 @@ def lift (s : cone F) : /-- Construct limit data for a product in `AddCommGroup`, using `AddCommGroup.of (Π j, F.obj j)`. -/ -def product_limit_cone : limits.limit_cone F := +@[simps] def product_limit_cone : limits.limit_cone F := { cone := { X := AddCommGroup.of (Π j, F.obj j), π := discrete.nat_trans (λ j, pi.eval_add_monoid_hom (λ j, F.obj j) j), }, @@ -92,29 +102,22 @@ def product_limit_cone : limits.limit_cone F := end has_limit -section - open has_limit -variables [decidable_eq J] [fintype J] - -instance (f : J → AddCommGroup.{u}) : has_biproduct f := -has_biproduct.of_has_product _ - /-- We verify that the biproduct we've just defined is isomorphic to the AddCommGroup structure on the dependent function type -/ -noncomputable -def biproduct_iso_pi (f : J → AddCommGroup.{u}) : +@[simps hom_apply] noncomputable +def biproduct_iso_pi [decidable_eq J] [fintype J] (f : J → AddCommGroup.{u}) : (⨁ f : AddCommGroup) ≅ AddCommGroup.of (Π j, f j) := is_limit.cone_point_unique_up_to_iso (biproduct.is_limit f) (product_limit_cone (discrete.functor f)).is_limit -end - -instance : has_finite_biproducts AddCommGroup := -⟨λ J _ _, by exactI { has_biproduct := λ f, by apply_instance }⟩ +@[simp, elementwise] lemma biproduct_iso_pi_inv_comp_π [decidable_eq J] [fintype J] + (f : J → AddCommGroup.{u}) (j : J) : + (biproduct_iso_pi f).inv ≫ biproduct.π f j = pi.eval_add_monoid_hom (λ j, f j) j := +is_limit.cone_point_unique_up_to_iso_inv_comp _ _ _ end AddCommGroup diff --git a/src/algebra/category/Module/biproducts.lean b/src/algebra/category/Module/biproducts.lean new file mode 100644 index 0000000000000..ee08a2841074a --- /dev/null +++ b/src/algebra/category/Module/biproducts.lean @@ -0,0 +1,129 @@ +/- +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.group.pi +import category_theory.limits.shapes.biproducts +import algebra.category.Module.limits + +/-! +# The category of `R`-modules has finite biproducts +-/ + +open category_theory +open category_theory.limits + +open_locale big_operators + +universes v u + +namespace Module + +variables {R : Type u} [ring R] + +-- As `Module R` is preadditive, and has all limits, it automatically has biproducts. +instance : has_binary_biproducts (Module.{v} R) := +has_binary_biproducts.of_has_binary_products + +instance : has_finite_biproducts (Module.{v} R) := +has_finite_biproducts.of_has_finite_products + +-- We now construct explicit limit data, +-- so we can compare the biproducts to the usual unbundled constructions. + +/-- +Construct limit data for a binary product in `Module R`, using `Module.of R (M × N)`. +-/ +@[simps cone_X is_limit_lift] +def binary_product_limit_cone (M N : Module.{v} R) : limits.limit_cone (pair M N) := +{ cone := + { X := Module.of R (M × N), + π := { app := λ j, walking_pair.cases_on j (linear_map.fst R M N) (linear_map.snd R M N) }}, + is_limit := + { lift := λ s, linear_map.prod (s.π.app walking_pair.left) (s.π.app walking_pair.right), + fac' := by { rintros s (⟨⟩|⟨⟩); { ext x, simp, }, }, + uniq' := λ s m w, + begin + ext; [rw ← w walking_pair.left, rw ← w walking_pair.right]; refl, + end, } } + +@[simp] lemma binary_product_limit_cone_cone_π_app_left (M N : Module.{v} R) : + (binary_product_limit_cone M N).cone.π.app walking_pair.left = linear_map.fst R M N := rfl + +@[simp] lemma binary_product_limit_cone_cone_π_app_right (M N : Module.{v} R) : + (binary_product_limit_cone M N).cone.π.app walking_pair.right = linear_map.snd R M N := rfl + +/-- +We verify that the biproduct in `Module R` is isomorphic to +the cartesian product of the underlying types: +-/ +@[simps hom_apply] noncomputable +def biprod_iso_prod (M N : Module.{v} R) : (M ⊞ N : Module.{v} R) ≅ Module.of R (M × N) := +is_limit.cone_point_unique_up_to_iso + (binary_biproduct.is_limit M N) + (binary_product_limit_cone M N).is_limit + +@[simp, elementwise] lemma biprod_iso_prod_inv_comp_fst (M N : Module.{v} R) : + (biprod_iso_prod M N).inv ≫ biprod.fst = linear_map.fst R M N := +is_limit.cone_point_unique_up_to_iso_inv_comp _ _ walking_pair.left + +@[simp, elementwise] lemma biprod_iso_prod_inv_comp_snd (M N : Module.{v} R) : + (biprod_iso_prod M N).inv ≫ biprod.snd = linear_map.snd R M N := +is_limit.cone_point_unique_up_to_iso_inv_comp _ _ walking_pair.right + +variables {J : Type v} (F : (discrete J) ⥤ Module.{v} R) + +namespace has_limit + +/-- +The map from an arbitrary cone over a indexed family of abelian groups +to the cartesian product of those groups. +-/ +def lift (s : cone F) : + s.X ⟶ Module.of R (Π j, F.obj j) := +{ to_fun := λ x j, s.π.app j x, + map_add' := λ x y, by { ext, simp, }, + map_smul' := λ r x, by { ext, simp, }, } + +@[simp] lemma lift_apply (s : cone F) (x : s.X) (j : J) : (lift F s) x j = s.π.app j x := rfl + +/-- +Construct limit data for a product in `Module R`, using `Module.of R (Π j, F.obj j)`. +-/ +@[simps] def product_limit_cone : limits.limit_cone F := +{ cone := + { X := Module.of R (Π j, F.obj j), + π := discrete.nat_trans (λ j, (linear_map.proj j : (Π j, F.obj j) →ₗ[R] F.obj j)), }, + is_limit := + { lift := lift F, + fac' := λ s j, by { ext, simp, }, + uniq' := λ s m w, + begin + ext x j, + dsimp only [has_limit.lift], + simp only [linear_map.coe_mk], + exact congr_arg (λ f : s.X ⟶ F.obj j, (f : s.X → F.obj j) x) (w j), + end, }, } + +end has_limit + +open has_limit + +/-- +We verify that the biproduct we've just defined is isomorphic to the `Module R` structure +on the dependent function type +-/ +@[simps hom_apply] noncomputable +def biproduct_iso_pi [decidable_eq J] [fintype J] (f : J → Module.{v} R) : + (⨁ f : Module.{v} R) ≅ Module.of R (Π j, f j) := +is_limit.cone_point_unique_up_to_iso + (biproduct.is_limit f) + (product_limit_cone (discrete.functor f)).is_limit + +@[simp, elementwise] lemma biproduct_iso_pi_inv_comp_π [decidable_eq J] [fintype J] + (f : J → Module.{v} R) (j : J) : + (biproduct_iso_pi f).inv ≫ biproduct.π f j = (linear_map.proj j : (Π j, f j) →ₗ[R] f j) := +is_limit.cone_point_unique_up_to_iso_inv_comp _ _ _ + +end Module From 3478a2af22c9eacc1b064f49514e7acea46fbeaa Mon Sep 17 00:00:00 2001 From: sgouezel Date: Sun, 8 May 2022 12:55:28 +0000 Subject: [PATCH 19/37] feat(probability/ident_distrib): identically distributed random variables (#14024) --- src/probability/ident_distrib.lean | 285 +++++++++++++++++++++++++++++ 1 file changed, 285 insertions(+) create mode 100644 src/probability/ident_distrib.lean diff --git a/src/probability/ident_distrib.lean b/src/probability/ident_distrib.lean new file mode 100644 index 0000000000000..2424ce81a4e8b --- /dev/null +++ b/src/probability/ident_distrib.lean @@ -0,0 +1,285 @@ +/- +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 probability.variance + +/-! +# Identically distributed random variables + +Two random variables defined on two (possibly different) probability spaces but taking value in +the same space are *identically distributed* if their distributions (i.e., the image probability +measures on the target space) coincide. We define this concept and establish its basic properties +in this file. + +## Main definitions and results + +* `ident_distrib f g μ ν` registers that the image of `μ` under `f` coincides with the image of `ν` + under `g` (and that `f` and `g` are almost everywhere measurable, as otherwise the image measures + don't make sense). The measures can be kept implicit as in `ident_distrib f g` if the spaces + are registered as measure spaces. +* `ident_distrib.comp`: being identically distributed is stable under composition with measurable + maps. + +There are two main kind of lemmas, under the assumption that `f` and `g` are identically +distributed: lemmas saying that two quantities computed for `f` and `g` are the same, and lemmas +saying that if `f` has some property then `g` also has it. The first kind is registered as +`ident_distrib.foo_eq`, the second one as `ident_distrib.foo_snd` (in the latter case, to deduce +a property of `f` from one of `g`, use `h.symm.foo_snd` where `h : ident_distrib f g μ ν`). For +instance: + +* `ident_distrib.measure_mem_eq`: if `f` and `g` are identically distributed, then the probabilities + that they belong to a given measurable set are the same. +* `ident_distrib.integral_eq`: if `f` and `g` are identically distributed, then their integrals + are the same. +* `ident_distrib.variance_eq`: if `f` and `g` are identically distributed, then their variances + are the same. + +* `ident_distrib.ae_strongly_measurable_snd`: if `f` and `g` are identically distributed and `f` + is almost everywhere strongly measurable, then so is `g`. +* `ident_distrib.mem_ℒp_snd`: if `f` and `g` are identically distributed and `f` + belongs to `ℒp`, then so does `g`. + +We also register several dot notation shortcuts for convenience. +For instance, if `h : ident_distrib f g μ ν`, then `h.sq` states that `f^2` and `g^2` are +identically distributed, and `h.norm` states that `∥f∥` and `∥g∥` are identically distributed, and +so on. +-/ + +open measure_theory filter finset + +noncomputable theory + +open_locale topological_space big_operators measure_theory ennreal nnreal + +variables {α β γ δ : Type*} [measurable_space α] [measurable_space β] + [measurable_space γ] [measurable_space δ] + +namespace probability_theory + +/-- Two functions defined on two (possibly different) measure spaces are identically distributed if +their image measures coincide. This only makes sense when the functions are ae measurable +(as otherwise the image measures are not defined), so we require this as well in the definition. -/ +structure ident_distrib + (f : α → γ) (g : β → γ) (μ : measure α . volume_tac) (ν : measure β . volume_tac) : Prop := +(ae_measurable_fst : ae_measurable f μ) +(ae_measurable_snd : ae_measurable g ν) +(map_eq : measure.map f μ = measure.map g ν) + +namespace ident_distrib + +open topological_space + +variables {μ : measure α} {ν : measure β} {f : α → γ} {g : β → γ} + +protected lemma refl (hf : ae_measurable f μ) : + ident_distrib f f μ μ := +{ ae_measurable_fst := hf, + ae_measurable_snd := hf, + map_eq := rfl } + +protected lemma symm (h : ident_distrib f g μ ν) : ident_distrib g f ν μ := +{ ae_measurable_fst := h.ae_measurable_snd, + ae_measurable_snd := h.ae_measurable_fst, + map_eq := h.map_eq.symm } + +protected lemma trans {ρ : measure δ} {h : δ → γ} + (h₁ : ident_distrib f g μ ν) (h₂ : ident_distrib g h ν ρ) : ident_distrib f h μ ρ := +{ ae_measurable_fst := h₁.ae_measurable_fst, + ae_measurable_snd := h₂.ae_measurable_snd, + map_eq := h₁.map_eq.trans h₂.map_eq } + +protected lemma comp_of_ae_measurable {u : γ → δ} (h : ident_distrib f g μ ν) + (hu : ae_measurable u (measure.map f μ)) : + ident_distrib (u ∘ f) (u ∘ g) μ ν := +{ ae_measurable_fst := hu.comp_ae_measurable h.ae_measurable_fst, + ae_measurable_snd := + by { rw h.map_eq at hu, exact hu.comp_ae_measurable h.ae_measurable_snd }, + map_eq := + begin + rw [← ae_measurable.map_map_of_ae_measurable hu h.ae_measurable_fst, + ← ae_measurable.map_map_of_ae_measurable _ h.ae_measurable_snd, h.map_eq], + rwa ← h.map_eq, + end } + +protected lemma comp {u : γ → δ} (h : ident_distrib f g μ ν) (hu : measurable u) : + ident_distrib (u ∘ f) (u ∘ g) μ ν := +h.comp_of_ae_measurable hu.ae_measurable + +lemma measure_mem_eq (h : ident_distrib f g μ ν) {s : set γ} (hs : measurable_set s) : + μ (f ⁻¹' s) = ν (g ⁻¹' s) := +by rw [← measure.map_apply_of_ae_measurable h.ae_measurable_fst hs, + ← measure.map_apply_of_ae_measurable h.ae_measurable_snd hs, h.map_eq] + +alias measure_mem_eq ← probability_theory.ident_distrib.measure_preimage_eq + +lemma ae_snd (h : ident_distrib f g μ ν) {p : γ → Prop} + (pmeas : measurable_set {x | p x}) (hp : ∀ᵐ x ∂μ, p (f x)) : + ∀ᵐ x ∂ν, p (g x) := +begin + apply (ae_map_iff h.ae_measurable_snd pmeas).1, + rw ← h.map_eq, + exact (ae_map_iff h.ae_measurable_fst pmeas).2 hp, +end + +lemma ae_mem_snd (h : ident_distrib f g μ ν) {t : set γ} + (tmeas : measurable_set t) (ht : ∀ᵐ x ∂μ, f x ∈ t) : + ∀ᵐ x ∂ν, g x ∈ t := +h.ae_snd tmeas ht + +/-- In a second countable topology, the first function in an identically distributed pair is a.e. +strongly measurable. So is the second function, but use `h.symm.ae_strongly_measurable_fst` as +`h.ae_strongly_measurable_snd` has a different meaning.-/ +lemma ae_strongly_measurable_fst [topological_space γ] + [metrizable_space γ] [opens_measurable_space γ] [second_countable_topology γ] + (h : ident_distrib f g μ ν) : + ae_strongly_measurable f μ := +h.ae_measurable_fst.ae_strongly_measurable + +/-- If `f` and `g` are identically distributed and `f` is a.e. strongly measurable, so is `g`. -/ +lemma ae_strongly_measurable_snd [topological_space γ] [metrizable_space γ] [borel_space γ] + (h : ident_distrib f g μ ν) (hf : ae_strongly_measurable f μ) : + ae_strongly_measurable g ν := +begin + refine ae_strongly_measurable_iff_ae_measurable_separable.2 ⟨h.ae_measurable_snd, _⟩, + rcases (ae_strongly_measurable_iff_ae_measurable_separable.1 hf).2 with ⟨t, t_sep, ht⟩, + refine ⟨closure t, t_sep.closure, _⟩, + apply h.ae_mem_snd is_closed_closure.measurable_set, + filter_upwards [ht] with x hx using subset_closure hx, +end + +lemma ae_strongly_measurable_iff [topological_space γ] [metrizable_space γ] [borel_space γ] + (h : ident_distrib f g μ ν) : + ae_strongly_measurable f μ ↔ ae_strongly_measurable g ν := +⟨λ hf, h.ae_strongly_measurable_snd hf, λ hg, h.symm.ae_strongly_measurable_snd hg⟩ + +lemma ess_sup_eq [conditionally_complete_linear_order γ] [topological_space γ] + [opens_measurable_space γ] [order_closed_topology γ] (h : ident_distrib f g μ ν) : + ess_sup f μ = ess_sup g ν := +begin + have I : ∀ a, μ {x : α | a < f x} = ν {x : β | a < g x} := + λ a, h.measure_mem_eq measurable_set_Ioi, + simp_rw [ess_sup_eq_Inf, I], +end + +lemma lintegral_eq {f : α → ℝ≥0∞} {g : β → ℝ≥0∞} (h : ident_distrib f g μ ν) : + ∫⁻ x, f x ∂μ = ∫⁻ x, g x ∂ν := +begin + change ∫⁻ x, id (f x) ∂μ = ∫⁻ x, id (g x) ∂ν, + rw [← lintegral_map' ae_measurable_id h.ae_measurable_fst, + ← lintegral_map' ae_measurable_id h.ae_measurable_snd, h.map_eq], +end + +lemma integral_eq [normed_group γ] [normed_space ℝ γ] [complete_space γ] [borel_space γ] + (h : ident_distrib f g μ ν) : ∫ x, f x ∂μ = ∫ x, g x ∂ν := +begin + by_cases hf : ae_strongly_measurable f μ, + { have A : ae_strongly_measurable id (measure.map f μ), + { rw ae_strongly_measurable_iff_ae_measurable_separable, + rcases (ae_strongly_measurable_iff_ae_measurable_separable.1 hf).2 with ⟨t, t_sep, ht⟩, + refine ⟨ae_measurable_id, ⟨closure t, t_sep.closure, _⟩⟩, + rw ae_map_iff h.ae_measurable_fst, + { filter_upwards [ht] with x hx using subset_closure hx }, + { exact is_closed_closure.measurable_set } }, + change ∫ x, id (f x) ∂μ = ∫ x, id (g x) ∂ν, + rw [← integral_map h.ae_measurable_fst A], + rw h.map_eq at A, + rw [← integral_map h.ae_measurable_snd A, h.map_eq] }, + { rw integral_non_ae_strongly_measurable hf, + rw h.ae_strongly_measurable_iff at hf, + rw integral_non_ae_strongly_measurable hf } +end + +lemma snorm_eq [normed_group γ] [opens_measurable_space γ] (h : ident_distrib f g μ ν) (p : ℝ≥0∞) : + snorm f p μ = snorm g p ν := +begin + by_cases h0 : p = 0, + { simp [h0], }, + by_cases h_top : p = ∞, + { simp only [h_top, snorm, snorm_ess_sup, ennreal.top_ne_zero, eq_self_iff_true, if_true, + if_false], + apply ess_sup_eq, + exact h.comp (measurable_coe_nnreal_ennreal.comp measurable_nnnorm) }, + simp only [snorm_eq_snorm' h0 h_top, snorm', one_div], + congr' 1, + apply lintegral_eq, + exact h.comp + (measurable.pow_const (measurable_coe_nnreal_ennreal.comp measurable_nnnorm) p.to_real), +end + +lemma mem_ℒp_snd [normed_group γ] [borel_space γ] + {p : ℝ≥0∞} (h : ident_distrib f g μ ν) (hf : mem_ℒp f p μ) : + mem_ℒp g p ν := +begin + refine ⟨h.ae_strongly_measurable_snd hf.ae_strongly_measurable, _⟩, + rw ← h.snorm_eq, + exact hf.2 +end + +lemma mem_ℒp_iff [normed_group γ] [borel_space γ] {p : ℝ≥0∞} (h : ident_distrib f g μ ν) : + mem_ℒp f p μ ↔ mem_ℒp g p ν := +⟨λ hf, h.mem_ℒp_snd hf, λ hg, h.symm.mem_ℒp_snd hg⟩ + +lemma integrable_snd [normed_group γ] [borel_space γ] (h : ident_distrib f g μ ν) + (hf : integrable f μ) : integrable g ν := +begin + rw ← mem_ℒp_one_iff_integrable at hf ⊢, + exact h.mem_ℒp_snd hf +end + +lemma integrable_iff [normed_group γ] [borel_space γ] (h : ident_distrib f g μ ν) : + integrable f μ ↔ integrable g ν := +⟨λ hf, h.integrable_snd hf, λ hg, h.symm.integrable_snd hg⟩ + +protected lemma norm [normed_group γ] [borel_space γ] (h : ident_distrib f g μ ν) : + ident_distrib (λ x, ∥f x∥) (λ x, ∥g x∥) μ ν := +h.comp measurable_norm + +protected lemma nnnorm [normed_group γ] [borel_space γ] (h : ident_distrib f g μ ν) : + ident_distrib (λ x, ∥f x∥₊) (λ x, ∥g x∥₊) μ ν := +h.comp measurable_nnnorm + +protected lemma pow [has_pow γ ℕ] [has_measurable_pow γ ℕ] (h : ident_distrib f g μ ν) {n : ℕ} : + ident_distrib (λ x, (f x) ^ n) (λ x, (g x) ^ n) μ ν := +h.comp (measurable_id.pow_const n) + +protected lemma sq [has_pow γ ℕ] [has_measurable_pow γ ℕ] (h : ident_distrib f g μ ν) : + ident_distrib (λ x, (f x) ^ 2) (λ x, (g x) ^ 2) μ ν := +h.comp (measurable_id.pow_const 2) + +protected lemma coe_nnreal_ennreal {f : α → ℝ≥0} {g : β → ℝ≥0} (h : ident_distrib f g μ ν) : + ident_distrib (λ x, (f x : ℝ≥0∞)) (λ x, (g x : ℝ≥0∞)) μ ν := +h.comp measurable_coe_nnreal_ennreal + +@[to_additive] +lemma mul_const [has_mul γ] [has_measurable_mul γ] (h : ident_distrib f g μ ν) (c : γ) : + ident_distrib (λ x, f x * c) (λ x, g x * c) μ ν := +h.comp (measurable_mul_const c) + +@[to_additive] +lemma const_mul [has_mul γ] [has_measurable_mul γ] (h : ident_distrib f g μ ν) (c : γ) : + ident_distrib (λ x, c * f x) (λ x, c * g x) μ ν := +h.comp (measurable_const_mul c) + +@[to_additive] +lemma div_const [has_div γ] [has_measurable_div γ] (h : ident_distrib f g μ ν) (c : γ) : + ident_distrib (λ x, f x / c) (λ x, g x / c) μ ν := +h.comp (has_measurable_div.measurable_div_const c) + +@[to_additive] +lemma const_div [has_div γ] [has_measurable_div γ] (h : ident_distrib f g μ ν) (c : γ) : + 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 ν := +begin + convert (h.sub_const (∫ x, f x ∂μ)).sq.integral_eq, + rw h.integral_eq, + refl +end + +end ident_distrib + +end probability_theory From 79ffb5563b56fefdea3d60b5736dad168a9494ab Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sun, 8 May 2022 14:02:33 +0000 Subject: [PATCH 20/37] chore(algebra/category/CommRing=>Ring): rename (#14022) This folder was originally named `algebra/category/CommRing/` because it only handled the commutative case. That's largely no longer the case, so we should rename the folder. Co-authored-by: Scott Morrison --- src/algebra/category/Algebra/basic.lean | 2 +- src/algebra/category/Algebra/limits.lean | 2 +- src/algebra/category/BoolRing.lean | 2 +- src/algebra/category/CommRing/default.lean | 3 --- src/algebra/category/{CommRing => Ring}/adjunctions.lean | 2 +- src/algebra/category/{CommRing => Ring}/basic.lean | 0 src/algebra/category/{CommRing => Ring}/colimits.lean | 2 +- src/algebra/category/{CommRing => Ring}/constructions.lean | 4 ++-- src/algebra/category/Ring/default.lean | 3 +++ .../category/{CommRing => Ring}/filtered_colimits.lean | 2 +- src/algebra/category/{CommRing => Ring}/instances.lean | 2 +- src/algebra/category/{CommRing => Ring}/limits.lean | 2 +- .../locally_ringed_space/has_colimits.lean | 2 +- src/algebraic_geometry/open_immersion.lean | 2 +- src/algebraic_geometry/properties.lean | 2 +- src/algebraic_geometry/ringed_space.lean | 6 +++--- src/algebraic_geometry/structure_sheaf.lean | 4 ++-- src/ring_theory/ideal/local_ring.lean | 2 +- src/topology/category/TopCommRing.lean | 2 +- src/topology/sheaves/forget.lean | 2 +- src/topology/sheaves/stalks.lean | 2 +- 21 files changed, 25 insertions(+), 25 deletions(-) delete mode 100644 src/algebra/category/CommRing/default.lean rename src/algebra/category/{CommRing => Ring}/adjunctions.lean (97%) rename src/algebra/category/{CommRing => Ring}/basic.lean (100%) rename src/algebra/category/{CommRing => Ring}/colimits.lean (99%) rename src/algebra/category/{CommRing => Ring}/constructions.lean (98%) create mode 100644 src/algebra/category/Ring/default.lean rename src/algebra/category/{CommRing => Ring}/filtered_colimits.lean (99%) rename src/algebra/category/{CommRing => Ring}/instances.lean (94%) rename src/algebra/category/{CommRing => Ring}/limits.lean (99%) diff --git a/src/algebra/category/Algebra/basic.lean b/src/algebra/category/Algebra/basic.lean index 8d7e5cd952b63..b0630ea5f9992 100644 --- a/src/algebra/category/Algebra/basic.lean +++ b/src/algebra/category/Algebra/basic.lean @@ -5,7 +5,7 @@ Authors: Scott Morrison -/ import algebra.algebra.subalgebra.basic import algebra.free_algebra -import algebra.category.CommRing.basic +import algebra.category.Ring.basic import algebra.category.Module.basic /-! diff --git a/src/algebra/category/Algebra/limits.lean b/src/algebra/category/Algebra/limits.lean index bfed0f4c8193e..dae9981edb01e 100644 --- a/src/algebra/category/Algebra/limits.lean +++ b/src/algebra/category/Algebra/limits.lean @@ -5,7 +5,7 @@ Authors: Scott Morrison -/ import algebra.category.Algebra.basic import algebra.category.Module.limits -import algebra.category.CommRing.limits +import algebra.category.Ring.limits /-! # The category of R-algebras has all limits diff --git a/src/algebra/category/BoolRing.lean b/src/algebra/category/BoolRing.lean index da745316e414d..87b85254a8bf9 100644 --- a/src/algebra/category/BoolRing.lean +++ b/src/algebra/category/BoolRing.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 algebra.category.CommRing.basic +import algebra.category.Ring.basic import algebra.ring.boolean_ring import order.category.BoolAlg diff --git a/src/algebra/category/CommRing/default.lean b/src/algebra/category/CommRing/default.lean deleted file mode 100644 index d2ebccabca5c2..0000000000000 --- a/src/algebra/category/CommRing/default.lean +++ /dev/null @@ -1,3 +0,0 @@ -import algebra.category.CommRing.adjunctions -import algebra.category.CommRing.limits -import algebra.category.CommRing.colimits diff --git a/src/algebra/category/CommRing/adjunctions.lean b/src/algebra/category/Ring/adjunctions.lean similarity index 97% rename from src/algebra/category/CommRing/adjunctions.lean rename to src/algebra/category/Ring/adjunctions.lean index 9cb12d59e49ca..d1d6809d85fea 100644 --- a/src/algebra/category/CommRing/adjunctions.lean +++ b/src/algebra/category/Ring/adjunctions.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, Johannes Hölzl -/ -import algebra.category.CommRing.basic +import algebra.category.Ring.basic import data.mv_polynomial.comm_ring /-! diff --git a/src/algebra/category/CommRing/basic.lean b/src/algebra/category/Ring/basic.lean similarity index 100% rename from src/algebra/category/CommRing/basic.lean rename to src/algebra/category/Ring/basic.lean diff --git a/src/algebra/category/CommRing/colimits.lean b/src/algebra/category/Ring/colimits.lean similarity index 99% rename from src/algebra/category/CommRing/colimits.lean rename to src/algebra/category/Ring/colimits.lean index f4a3b8d687c1f..f786d20f28eb3 100644 --- a/src/algebra/category/CommRing/colimits.lean +++ b/src/algebra/category/Ring/colimits.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 algebra.category.CommRing.basic +import algebra.category.Ring.basic import category_theory.limits.has_limits import category_theory.limits.concrete_category diff --git a/src/algebra/category/CommRing/constructions.lean b/src/algebra/category/Ring/constructions.lean similarity index 98% rename from src/algebra/category/CommRing/constructions.lean rename to src/algebra/category/Ring/constructions.lean index 47217784f4c03..9e9dc9e24b503 100644 --- a/src/algebra/category/CommRing/constructions.lean +++ b/src/algebra/category/Ring/constructions.lean @@ -5,8 +5,8 @@ Authors: Andrew Yang -/ import category_theory.limits.shapes.pullbacks import ring_theory.tensor_product -import algebra.category.CommRing.limits -import algebra.category.CommRing.colimits +import algebra.category.Ring.limits +import algebra.category.Ring.colimits import category_theory.limits.shapes.strict_initial import ring_theory.subring.basic import ring_theory.ideal.local_ring diff --git a/src/algebra/category/Ring/default.lean b/src/algebra/category/Ring/default.lean new file mode 100644 index 0000000000000..ef2d0c7f26ad5 --- /dev/null +++ b/src/algebra/category/Ring/default.lean @@ -0,0 +1,3 @@ +import algebra.category.Ring.adjunctions +import algebra.category.Ring.limits +import algebra.category.Ring.colimits diff --git a/src/algebra/category/CommRing/filtered_colimits.lean b/src/algebra/category/Ring/filtered_colimits.lean similarity index 99% rename from src/algebra/category/CommRing/filtered_colimits.lean rename to src/algebra/category/Ring/filtered_colimits.lean index 34406fe037c8d..25ca46a650d8f 100644 --- a/src/algebra/category/CommRing/filtered_colimits.lean +++ b/src/algebra/category/Ring/filtered_colimits.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Justus Springer. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Justus Springer -/ -import algebra.category.CommRing.basic +import algebra.category.Ring.basic import algebra.category.Group.filtered_colimits /-! diff --git a/src/algebra/category/CommRing/instances.lean b/src/algebra/category/Ring/instances.lean similarity index 94% rename from src/algebra/category/CommRing/instances.lean rename to src/algebra/category/Ring/instances.lean index 302566bfbc8b6..55c663a16477d 100644 --- a/src/algebra/category/CommRing/instances.lean +++ b/src/algebra/category/Ring/instances.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Andrew Yang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ -import algebra.category.CommRing.basic +import algebra.category.Ring.basic import ring_theory.localization.away /-! diff --git a/src/algebra/category/CommRing/limits.lean b/src/algebra/category/Ring/limits.lean similarity index 99% rename from src/algebra/category/CommRing/limits.lean rename to src/algebra/category/Ring/limits.lean index f3451621813f9..0de84e5199c3e 100644 --- a/src/algebra/category/CommRing/limits.lean +++ b/src/algebra/category/Ring/limits.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ import algebra.ring.pi -import algebra.category.CommRing.basic +import algebra.category.Ring.basic import algebra.category.Group.limits import ring_theory.subring.basic diff --git a/src/algebraic_geometry/locally_ringed_space/has_colimits.lean b/src/algebraic_geometry/locally_ringed_space/has_colimits.lean index 914c360cfb4d0..fc14a688ec35a 100644 --- a/src/algebraic_geometry/locally_ringed_space/has_colimits.lean +++ b/src/algebraic_geometry/locally_ringed_space/has_colimits.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ import algebraic_geometry.locally_ringed_space -import algebra.category.CommRing.constructions +import algebra.category.Ring.constructions import algebraic_geometry.open_immersion import category_theory.limits.constructions.limits_of_products_and_equalizers diff --git a/src/algebraic_geometry/open_immersion.lean b/src/algebraic_geometry/open_immersion.lean index cfd68e0f51f0d..3c82312777987 100644 --- a/src/algebraic_geometry/open_immersion.lean +++ b/src/algebraic_geometry/open_immersion.lean @@ -9,7 +9,7 @@ import category_theory.limits.preserves.shapes.pullbacks import topology.sheaves.functors import algebraic_geometry.Scheme import category_theory.limits.shapes.strict_initial -import algebra.category.CommRing.instances +import algebra.category.Ring.instances /-! # Open immersions of structured spaces diff --git a/src/algebraic_geometry/properties.lean b/src/algebraic_geometry/properties.lean index 3bdcd876ef3bf..533a7c50b2d20 100644 --- a/src/algebraic_geometry/properties.lean +++ b/src/algebraic_geometry/properties.lean @@ -7,7 +7,7 @@ import algebraic_geometry.AffineScheme import ring_theory.nilpotent import topology.sheaves.sheaf_condition.sites import category_theory.limits.constructions.binary_products -import algebra.category.CommRing.constructions +import algebra.category.Ring.constructions import ring_theory.integral_domain import ring_theory.local_properties diff --git a/src/algebraic_geometry/ringed_space.lean b/src/algebraic_geometry/ringed_space.lean index fc0c6a260b7eb..97ef0e5241d69 100644 --- a/src/algebraic_geometry/ringed_space.lean +++ b/src/algebraic_geometry/ringed_space.lean @@ -3,11 +3,11 @@ Copyright (c) 2021 Justus Springer. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Justus Springer, Andrew Yang -/ -import algebra.category.CommRing.filtered_colimits +import algebra.category.Ring.filtered_colimits import algebraic_geometry.sheafed_space import topology.sheaves.stalks -import algebra.category.CommRing.colimits -import algebra.category.CommRing.limits +import algebra.category.Ring.colimits +import algebra.category.Ring.limits /-! # Ringed spaces diff --git a/src/algebraic_geometry/structure_sheaf.lean b/src/algebraic_geometry/structure_sheaf.lean index f80c5316ccc7a..efbf261969500 100644 --- a/src/algebraic_geometry/structure_sheaf.lean +++ b/src/algebraic_geometry/structure_sheaf.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Scott Morrison -/ import algebraic_geometry.prime_spectrum.basic -import algebra.category.CommRing.colimits -import algebra.category.CommRing.limits +import algebra.category.Ring.colimits +import algebra.category.Ring.limits import topology.sheaves.local_predicate import ring_theory.localization.at_prime import ring_theory.subring.basic diff --git a/src/ring_theory/ideal/local_ring.lean b/src/ring_theory/ideal/local_ring.lean index 77394649e09d7..1cd68f3c33f37 100644 --- a/src/ring_theory/ideal/local_ring.lean +++ b/src/ring_theory/ideal/local_ring.lean @@ -5,7 +5,7 @@ Authors: Kenny Lau, Chris Hughes, Mario Carneiro -/ import algebra.algebra.basic -import algebra.category.CommRing.basic +import algebra.category.Ring.basic import ring_theory.ideal.operations /-! diff --git a/src/topology/category/TopCommRing.lean b/src/topology/category/TopCommRing.lean index e0149c0f31aff..edddbbd6b5fd4 100644 --- a/src/topology/category/TopCommRing.lean +++ b/src/topology/category/TopCommRing.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 algebra.category.CommRing.basic +import algebra.category.Ring.basic import topology.category.Top.basic import topology.algebra.ring diff --git a/src/topology/sheaves/forget.lean b/src/topology/sheaves/forget.lean index 255145043d058..a90ac66653eba 100644 --- a/src/topology/sheaves/forget.lean +++ b/src/topology/sheaves/forget.lean @@ -201,7 +201,7 @@ As an example, we now have everything we need to check the sheaf condition for a presheaf of commutative rings, merely by checking the sheaf condition for the underlying sheaf of types. ``` -import algebra.category.CommRing.limits +import algebra.category.Ring.limits example (X : Top) (F : presheaf CommRing X) (h : presheaf.is_sheaf (F ⋙ (forget CommRing))) : F.is_sheaf := (is_sheaf_iff_is_sheaf_comp (forget CommRing) F).mpr h diff --git a/src/topology/sheaves/stalks.lean b/src/topology/sheaves/stalks.lean index 677705b23e3a7..620748c163de1 100644 --- a/src/topology/sheaves/stalks.lean +++ b/src/topology/sheaves/stalks.lean @@ -11,7 +11,7 @@ import category_theory.limits.preserves.filtered import category_theory.limits.final import topology.sober import tactic.elementwise -import algebra.category.CommRing +import algebra.category.Ring /-! # Stalks From e3306942ebcd15ba75ff24fb43ca2effa35ae8f8 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Sun, 8 May 2022 15:55:36 +0000 Subject: [PATCH 21/37] feat(analysis/p_series): explicit bounds on sums of the form 1/j^2 (#13851) --- src/analysis/p_series.lean | 60 ++++++++++++++++++++++++++++++++++++++ 1 file changed, 60 insertions(+) diff --git a/src/analysis/p_series.lean b/src/analysis/p_series.lean index 8abfa7279cfee..28bd743d1a821 100644 --- a/src/analysis/p_series.lean +++ b/src/analysis/p_series.lean @@ -229,3 +229,63 @@ by simp [← nnreal.summable_coe] lemma nnreal.summable_one_div_rpow {p : ℝ} : summable (λ n, 1 / n ^ p : ℕ → ℝ≥0) ↔ 1 < p := by simp + +section + +open finset + +variables {α : Type*} [linear_ordered_field α] + +lemma sum_Ioc_inv_sq_le_sub {k n : ℕ} (hk : k ≠ 0) (h : k ≤ n) : + ∑ i in Ioc k n, ((i ^ 2) ⁻¹ : α) ≤ k ⁻¹ - n ⁻¹ := +begin + refine nat.le_induction _ _ n h, + { simp only [Ioc_self, sum_empty, sub_self] }, + assume n hn IH, + rw [sum_Ioc_succ_top hn], + apply (add_le_add IH le_rfl).trans, + simp only [sub_eq_add_neg, add_assoc, nat.cast_add, nat.cast_one, le_add_neg_iff_add_le, + add_le_iff_nonpos_right, neg_add_le_iff_le_add, add_zero], + have A : 0 < (n : α), by simpa using hk.bot_lt.trans_le hn, + have B : 0 < (n : α) + 1, by linarith, + field_simp [B.ne'], + rw [div_le_div_iff _ A, ← sub_nonneg], + { ring_nf, exact B.le }, + { nlinarith }, +end + +lemma sum_Ioo_inv_sq_le (k n : ℕ) : + ∑ i in Ioo k n, ((i ^ 2) ⁻¹ : α) ≤ 2 / (k + 1) := +calc +∑ i in Ioo k n, ((i ^ 2) ⁻¹ : α) ≤ ∑ i in Ioc k (max (k+1) n), (i ^ 2) ⁻¹ : +begin + apply sum_le_sum_of_subset_of_nonneg, + { assume x hx, + simp only [mem_Ioo] at hx, + simp only [hx, hx.2.le, mem_Ioc, le_max_iff, or_true, and_self] }, + { assume i hi hident, + exact inv_nonneg.2 (sq_nonneg _), } +end +... ≤ ((k + 1) ^ 2) ⁻¹ + ∑ i in Ioc k.succ (max (k + 1) n), (i ^ 2) ⁻¹ : +begin + rw [← nat.Icc_succ_left, ← nat.Ico_succ_right, sum_eq_sum_Ico_succ_bot], + swap, { exact nat.succ_lt_succ ((nat.lt_succ_self k).trans_le (le_max_left _ _)) }, + rw [nat.Ico_succ_right, nat.Icc_succ_left, nat.cast_succ], +end +... ≤ ((k + 1) ^ 2) ⁻¹ + (k + 1) ⁻¹ : +begin + refine add_le_add le_rfl ((sum_Ioc_inv_sq_le_sub _ (le_max_left _ _)).trans _), + { simp only [ne.def, nat.succ_ne_zero, not_false_iff] }, + { simp only [nat.cast_succ, one_div, sub_le_self_iff, inv_nonneg, nat.cast_nonneg] } +end +... ≤ 1 / (k + 1) + 1 / (k + 1) : +begin + have A : (1 : α) ≤ k + 1, by simp only [le_add_iff_nonneg_left, nat.cast_nonneg], + simp_rw ← one_div, + apply add_le_add_right, + refine div_le_div zero_le_one le_rfl (zero_lt_one.trans_le A) _, + simpa using pow_le_pow A one_le_two, +end +... = 2 / (k + 1) : by ring + +end From 6a5d17eea69ac0948920944b1ff3bd0ddcaf999d Mon Sep 17 00:00:00 2001 From: sgouezel Date: Sun, 8 May 2022 16:31:28 +0000 Subject: [PATCH 22/37] feat(measure_theory/integral): a few more integral lemmas (#14025) --- .../integral/interval_integral.lean | 40 ++++++++++++++----- src/measure_theory/integral/set_integral.lean | 18 +++++++++ 2 files changed, 48 insertions(+), 10 deletions(-) diff --git a/src/measure_theory/integral/interval_integral.lean b/src/measure_theory/integral/interval_integral.lean index af85dc5380c35..9b527eac71164 100644 --- a/src/measure_theory/integral/interval_integral.lean +++ b/src/measure_theory/integral/interval_integral.lean @@ -238,14 +238,21 @@ by split; simp ⟨(hab.1.union hbc.1).mono_set Ioc_subset_Ioc_union_Ioc, (hbc.2.union hab.2).mono_set Ioc_subset_Ioc_union_Ioc⟩ -lemma trans_iterate {a : ℕ → ℝ} {n : ℕ} (hint : ∀ k < n, interval_integrable f μ (a k) (a $ k+1)) : - interval_integrable f μ (a 0) (a n) := +lemma trans_iterate_Ico {a : ℕ → ℝ} {m n : ℕ} (hmn : m ≤ n) + (hint : ∀ k ∈ Ico m n, interval_integrable f μ (a k) (a $ k+1)) : + interval_integrable f μ (a m) (a n) := begin - induction n with n hn, + revert hint, + refine nat.le_induction _ _ n hmn, { simp }, - { exact (hn (λ k hk, hint k (hk.trans n.lt_succ_self))).trans (hint n n.lt_succ_self) } + { assume p hp IH h, + exact (IH (λ k hk, h k (Ico_subset_Ico_right p.le_succ hk))).trans (h p (by simp [hp])) } end +lemma trans_iterate {a : ℕ → ℝ} {n : ℕ} (hint : ∀ k < n, interval_integrable f μ (a k) (a $ k+1)) : + interval_integrable f μ (a 0) (a n) := +trans_iterate_Ico bot_le (λ k hk, hint k hk.2) + lemma neg (h : interval_integrable f μ a b) : interval_integrable (-f) μ a b := ⟨h.1.neg, h.2.neg⟩ @@ -769,16 +776,29 @@ lemma integral_add_adjacent_intervals (hab : interval_integrable f μ a b) ∫ x in a..b, f x ∂μ + ∫ x in b..c, f x ∂μ = ∫ x in a..c, f x ∂μ := by rw [← add_neg_eq_zero, ← integral_symm, integral_add_adjacent_intervals_cancel hab hbc] +lemma sum_integral_adjacent_intervals_Ico {a : ℕ → ℝ} {m n : ℕ} (hmn : m ≤ n) + (hint : ∀ k ∈ Ico m n, interval_integrable f μ (a k) (a $ k+1)) : + ∑ (k : ℕ) in finset.Ico m n, ∫ x in (a k)..(a $ k+1), f x ∂μ = ∫ x in (a m)..(a n), f x ∂μ := +begin + revert hint, + refine nat.le_induction _ _ n hmn, + { simp }, + { assume p hmp IH h, + rw [finset.sum_Ico_succ_top hmp, IH, integral_add_adjacent_intervals], + { apply interval_integrable.trans_iterate_Ico hmp (λ k hk, h k _), + exact (Ico_subset_Ico le_rfl (nat.le_succ _)) hk }, + { apply h, + simp [hmp] }, + { assume k hk, + exact h _ (Ico_subset_Ico_right p.le_succ hk) } } +end + lemma sum_integral_adjacent_intervals {a : ℕ → ℝ} {n : ℕ} (hint : ∀ k < n, interval_integrable f μ (a k) (a $ k+1)) : ∑ (k : ℕ) in finset.range n, ∫ x in (a k)..(a $ k+1), f x ∂μ = ∫ x in (a 0)..(a n), f x ∂μ := begin - induction n with n hn, - { simp }, - { rw [finset.sum_range_succ, hn (λ k hk, hint k (hk.trans n.lt_succ_self))], - exact integral_add_adjacent_intervals - (interval_integrable.trans_iterate $ λ k hk, hint k (hk.trans n.lt_succ_self)) - (hint n n.lt_succ_self) } + rw ← nat.Ico_zero_eq_range, + exact sum_integral_adjacent_intervals_Ico (zero_le n) (λ k hk, hint k hk.2), end lemma integral_interval_sub_left (hab : interval_integrable f μ a b) diff --git a/src/measure_theory/integral/set_integral.lean b/src/measure_theory/integral/set_integral.lean index b6d15435593f0..9b20f3ce4b6bf 100644 --- a/src/measure_theory/integral/set_integral.lean +++ b/src/measure_theory/integral/set_integral.lean @@ -141,6 +141,24 @@ begin ... = ∫ x in s, f x ∂μ : by simp end +lemma of_real_set_integral_one_of_measure_ne_top {α : Type*} {m : measurable_space α} + {μ : measure α} {s : set α} (hs : μ s ≠ ∞) : + ennreal.of_real (∫ x in s, (1 : ℝ) ∂μ) = μ s := +calc +ennreal.of_real (∫ x in s, (1 : ℝ) ∂μ) + = ennreal.of_real (∫ x in s, ∥(1 : ℝ)∥ ∂μ) : by simp only [norm_one] +... = ∫⁻ x in s, 1 ∂μ : +begin + rw of_real_integral_norm_eq_lintegral_nnnorm (integrable_on_const.2 (or.inr hs.lt_top)), + simp only [nnnorm_one, ennreal.coe_one], +end +... = μ s : set_lintegral_one _ + +lemma of_real_set_integral_one {α : Type*} {m : measurable_space α} (μ : measure α) + [is_finite_measure μ] (s : set α) : + ennreal.of_real (∫ x in s, (1 : ℝ) ∂μ) = μ s := +of_real_set_integral_one_of_measure_ne_top (measure_ne_top μ s) + lemma integral_piecewise [decidable_pred (∈ s)] (hs : measurable_set s) {f g : α → E} (hf : integrable_on f s μ) (hg : integrable_on g sᶜ μ) : ∫ x, s.piecewise f g x ∂μ = ∫ x in s, f x ∂μ + ∫ x in sᶜ, g x ∂μ := From 69c07a4b1cc7bf04ba6a71fadd7c4729bd224e52 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sun, 8 May 2022 16:31:29 +0000 Subject: [PATCH 23/37] feat(linear_algebra/linear_pmap): `mk_span_singleton` of the same point (#14029) One more lemma about `mk_span_singleton'` and slightly better lemma names. --- src/linear_algebra/dual.lean | 4 +--- src/linear_algebra/linear_pmap.lean | 22 ++++++++++++---------- 2 files changed, 13 insertions(+), 13 deletions(-) diff --git a/src/linear_algebra/dual.lean b/src/linear_algebra/dual.lean index 2528620e623b9..5ec63c6bb7c54 100644 --- a/src/linear_algebra/dual.lean +++ b/src/linear_algebra/dual.lean @@ -758,9 +758,7 @@ begin refine ne_zero_of_eq_one _, have h : f.comp (K ∙ x).subtype = (linear_pmap.mk_span_singleton x 1 hx).to_fun := classical.some_spec (linear_pmap.mk_span_singleton x (1 : K) hx).to_fun.exists_extend, - rw linear_map.ext_iff at h, - convert h ⟨x, submodule.mem_span_singleton_self x⟩, - exact (linear_pmap.mk_span_singleton_apply' K hx 1).symm, + exact (fun_like.congr_fun h _).trans (linear_pmap.mk_span_singleton_apply _ hx _), end end field diff --git a/src/linear_algebra/linear_pmap.lean b/src/linear_algebra/linear_pmap.lean index 51eeab7cd1ddb..d97de7ec18a05 100644 --- a/src/linear_algebra/linear_pmap.lean +++ b/src/linear_algebra/linear_pmap.lean @@ -94,7 +94,7 @@ noncomputable def mk_span_singleton' (x : E) (y : F) (H : ∀ c : R, c • x = 0 @[simp] lemma domain_mk_span_singleton (x : E) (y : F) (H : ∀ c : R, c • x = 0 → c • y = 0) : (mk_span_singleton' x y H).domain = R ∙ x := rfl -@[simp] lemma mk_span_singleton_apply (x : E) (y : F) (H : ∀ c : R, c • x = 0 → c • y = 0) +@[simp] lemma mk_span_singleton'_apply (x : E) (y : F) (H : ∀ c : R, c • x = 0 → c • y = 0) (c : R) (h) : mk_span_singleton' x y H ⟨c • x, h⟩ = c • y := begin @@ -105,6 +105,11 @@ begin apply classical.some_spec (mem_span_singleton.1 h), end +@[simp] lemma mk_span_singleton'_apply_self (x : E) (y : F) (H : ∀ c : R, c • x = 0 → c • y = 0) + (h) : + mk_span_singleton' x y H ⟨x, h⟩ = y := +by convert mk_span_singleton'_apply x y H 1 _; rwa one_smul + /-- The unique `linear_pmap` on `span R {x}` that sends a non-zero vector `x` to `y`. This version works for modules over division rings. -/ @[reducible] noncomputable def mk_span_singleton {K E F : Type*} [division_ring K] @@ -113,14 +118,11 @@ This version works for modules over division rings. -/ mk_span_singleton' x y $ λ c hc, (smul_eq_zero.1 hc).elim (λ hc, by rw [hc, zero_smul]) (λ hx', absurd hx' hx) -lemma mk_span_singleton_apply' (K : Type*) {E F : Type*} [division_ring K] - [add_comm_group E] [module K E] [add_comm_group F] [module K F] {x : E} (hx : x ≠ 0) (y : F): - (mk_span_singleton x y hx).to_fun - ⟨x, (submodule.mem_span_singleton_self x : x ∈ submodule.span K {x})⟩ = y := -begin - convert linear_pmap.mk_span_singleton_apply x y _ (1 : K) _; - simp [submodule.mem_span_singleton_self], -end +lemma mk_span_singleton_apply (K : Type*) {E F : Type*} [division_ring K] + [add_comm_group E] [module K E] [add_comm_group F] [module K F] {x : E} (hx : x ≠ 0) (y : F) : + mk_span_singleton x y hx + ⟨x, (submodule.mem_span_singleton_self x : x ∈ submodule.span K {x})⟩ = y := +linear_pmap.mk_span_singleton'_apply_self _ _ _ _ /-- Projection to the first coordinate as a `linear_pmap` -/ protected def fst (p : submodule R E) (p' : submodule R F) : linear_pmap R (E × F) E := @@ -315,7 +317,7 @@ f.sup (mk_span_singleton x y (λ h₀, hx $ h₀.symm ▸ f.domain.zero_mem)) $ f.sup_span_singleton x y hx ⟨x' + c • x, mem_sup.2 ⟨x', hx', _, mem_span_singleton.2 ⟨c, rfl⟩, rfl⟩⟩ = f ⟨x', hx'⟩ + c • y := begin - erw [sup_apply _ ⟨x', hx'⟩ ⟨c • x, _⟩, mk_span_singleton_apply], + erw [sup_apply _ ⟨x', hx'⟩ ⟨c • x, _⟩, mk_span_singleton'_apply], refl, exact mem_span_singleton.2 ⟨c, rfl⟩ end From dd16a836c9dfbbec9668e738f11f582c74473c2f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fr=C3=A9d=C3=A9ric=20Dupuis?= Date: Sun, 8 May 2022 18:01:58 +0000 Subject: [PATCH 24/37] fix(topology/algebra/module/weak_dual): fix namespace issue, add a few extra lemmas (#13407) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR fixes a namespace issue in `weak_dual`, to ensure lemmas with names like `eval_continuous` are appropriately namespaced. Also, lemmas about continuity of the evaluation map have been copied from `weak_bilin` to `weak_dual`. Co-authored-by: Frédéric Dupuis <31101893+dupuisf@users.noreply.github.com> --- src/analysis/locally_convex/polar.lean | 2 +- src/analysis/normed_space/weak_dual.lean | 4 +-- .../finite_measure_weak_convergence.lean | 2 +- src/topology/algebra/module/weak_dual.lean | 36 ++++++++++++++----- 4 files changed, 31 insertions(+), 13 deletions(-) diff --git a/src/analysis/locally_convex/polar.lean b/src/analysis/locally_convex/polar.lean index 0eac9560e8499..6c2d73ba3e9be 100644 --- a/src/analysis/locally_convex/polar.lean +++ b/src/analysis/locally_convex/polar.lean @@ -106,7 +106,7 @@ lemma polar_weak_closed (s : set E) : begin rw polar_eq_Inter, refine is_closed_Inter (λ x, is_closed_Inter (λ _, _)), - exact is_closed_le (eval_continuous B.flip x).norm continuous_const, + exact is_closed_le (weak_bilin.eval_continuous B.flip x).norm continuous_const, end end normed_ring diff --git a/src/analysis/normed_space/weak_dual.lean b/src/analysis/normed_space/weak_dual.lean index f405d40bf1078..472fd416adc9c 100644 --- a/src/analysis/normed_space/weak_dual.lean +++ b/src/analysis/normed_space/weak_dual.lean @@ -116,7 +116,7 @@ def to_weak_dual : dual 𝕜 E ≃ₗ[𝕜] weak_dual 𝕜 E := linear_equiv.ref to_weak_dual.injective.eq_iff theorem to_weak_dual_continuous : continuous (λ (x' : dual 𝕜 E), x'.to_weak_dual) := -continuous_of_continuous_eval _ $ λ z, (inclusion_in_double_dual 𝕜 E z).continuous +weak_bilin.continuous_of_continuous_eval _ $ λ z, (inclusion_in_double_dual 𝕜 E z).continuous /-- For a normed space `E`, according to `to_weak_dual_continuous` the "identity mapping" `dual 𝕜 E → weak_dual 𝕜 E` is continuous. This definition implements it as a continuous linear @@ -169,7 +169,7 @@ is used. -/ lemma is_closed_polar (s : set E) : is_closed (polar 𝕜 s) := begin simp only [polar_def, set_of_forall], - exact is_closed_bInter (λ x hx, is_closed_Iic.preimage (eval_continuous _ _).norm) + exact is_closed_bInter (λ x hx, is_closed_Iic.preimage (weak_bilin.eval_continuous _ _).norm) end variable {𝕜} diff --git a/src/measure_theory/measure/finite_measure_weak_convergence.lean b/src/measure_theory/measure/finite_measure_weak_convergence.lean index af7c74d2b256c..48b58a3c8b60a 100644 --- a/src/measure_theory/measure/finite_measure_weak_convergence.lean +++ b/src/measure_theory/measure/finite_measure_weak_convergence.lean @@ -304,7 +304,7 @@ continuous_induced_dom depends continuously on the measure. -/ lemma continuous_test_against_nn_eval (f : α →ᵇ ℝ≥0) : continuous (λ (μ : finite_measure α), μ.test_against_nn f) := -(by apply (eval_continuous _ _).comp to_weak_dual_bcnn_continuous : +(by apply (weak_bilin.eval_continuous _ _).comp to_weak_dual_bcnn_continuous : continuous ((λ φ : weak_dual ℝ≥0 (α →ᵇ ℝ≥0), φ f) ∘ to_weak_dual_bcnn)) lemma tendsto_iff_weak_star_tendsto {γ : Type*} {F : filter γ} diff --git a/src/topology/algebra/module/weak_dual.lean b/src/topology/algebra/module/weak_dual.lean index 10c9126b6706f..ba75f3f954916 100644 --- a/src/topology/algebra/module/weak_dual.lean +++ b/src/topology/algebra/module/weak_dual.lean @@ -72,15 +72,17 @@ nolint has_inhabited_instance unused_arguments] def weak_bilin [comm_semiring 𝕜] [add_comm_monoid E] [module 𝕜 E] [add_comm_monoid F] [module 𝕜 F] (B : E →ₗ[𝕜] F →ₗ[𝕜] 𝕜) := E +namespace weak_bilin + instance [comm_semiring 𝕜] [a : add_comm_group E] [module 𝕜 E] [add_comm_monoid F] [module 𝕜 F] (B : E →ₗ[𝕜] F →ₗ[𝕜] 𝕜) : add_comm_group (weak_bilin B) := a @[priority 100] -instance module_weak_bilin [comm_semiring 𝕜] [comm_semiring 𝕝] [add_comm_group E] [module 𝕜 E] +instance module' [comm_semiring 𝕜] [comm_semiring 𝕝] [add_comm_group E] [module 𝕜 E] [add_comm_group F] [module 𝕜 F] [m : module 𝕝 E] (B : E →ₗ[𝕜] F →ₗ[𝕜] 𝕜) : module 𝕝 (weak_bilin B) := m -instance scalar_tower_weak_bilin [comm_semiring 𝕜] [comm_semiring 𝕝] [add_comm_group E] [module 𝕜 E] +instance [comm_semiring 𝕜] [comm_semiring 𝕝] [add_comm_group E] [module 𝕜 E] [add_comm_group F] [module 𝕜 F] [has_scalar 𝕝 𝕜] [module 𝕝 E] [s : is_scalar_tower 𝕝 𝕜 E] (B : E →ₗ[𝕜] F →ₗ[𝕜] 𝕜) : is_scalar_tower 𝕝 𝕜 (weak_bilin B) := s @@ -105,13 +107,13 @@ lemma continuous_of_continuous_eval [topological_space α] {g : α → weak_bili continuous_induced_rng (continuous_pi_iff.mpr h) /-- The coercion `(λ x y, B x y) : E → (F → 𝕜)` is an embedding. -/ -lemma bilin_embedding {B : E →ₗ[𝕜] F →ₗ[𝕜] 𝕜} (hB : function.injective B) : +lemma embedding {B : E →ₗ[𝕜] F →ₗ[𝕜] 𝕜} (hB : function.injective B) : embedding (λ (x : weak_bilin B) y, B x y) := function.injective.embedding_induced $ linear_map.coe_injective.comp hB theorem tendsto_iff_forall_eval_tendsto {l : filter α} {f : α → (weak_bilin B)} {x : weak_bilin B} (hB : function.injective B) : tendsto f l (𝓝 x) ↔ ∀ y, tendsto (λ i, B (f i) y) l (𝓝 (B x y)) := -by rw [← tendsto_pi_nhds, embedding.tendsto_nhds_iff (bilin_embedding hB)] +by rw [← tendsto_pi_nhds, embedding.tendsto_nhds_iff (embedding hB)] /-- Addition in `weak_space B` is continuous. -/ instance [has_continuous_add 𝕜] : has_continuous_add (weak_bilin B) := @@ -155,6 +157,8 @@ instance [has_continuous_add 𝕜] : topological_add_group (weak_bilin B) := end ring +end weak_bilin + end weak_topology section weak_star_topology @@ -178,6 +182,8 @@ def weak_dual (𝕜 E) [comm_semiring 𝕜] [topological_space 𝕜] [has_contin [has_continuous_const_smul 𝕜 𝕜] [add_comm_monoid E] [module 𝕜 E] [topological_space E] := weak_bilin (top_dual_pairing 𝕜 E) +namespace weak_dual + instance : inhabited (weak_dual 𝕜 E) := continuous_linear_map.inhabited instance weak_dual.add_monoid_hom_class : add_monoid_hom_class (weak_dual 𝕜 E) E 𝕜 := @@ -203,22 +209,34 @@ continuous_linear_map.distrib_mul_action /-- If `𝕜` is a topological module over a semiring `R` and scalar multiplication commutes with the multiplication on `𝕜`, then `weak_dual 𝕜 E` is a module over `R`. -/ -instance weak_dual_module (R) [semiring R] [module R 𝕜] [smul_comm_class 𝕜 R 𝕜] +instance module' (R) [semiring R] [module R 𝕜] [smul_comm_class 𝕜 R 𝕜] [has_continuous_const_smul R 𝕜] : module R (weak_dual 𝕜 E) := continuous_linear_map.module instance (M) [monoid M] [distrib_mul_action M 𝕜] [smul_comm_class 𝕜 M 𝕜] [has_continuous_const_smul M 𝕜] : has_continuous_const_smul M (weak_dual 𝕜 E) := -⟨λ m, continuous_induced_rng $ (coe_fn_continuous (top_dual_pairing 𝕜 E)).const_smul m⟩ +⟨λ m, continuous_induced_rng $ (weak_bilin.coe_fn_continuous (top_dual_pairing 𝕜 E)).const_smul m⟩ /-- If a monoid `M` distributively continuously acts on `𝕜` and this action commutes with multiplication on `𝕜`, then it continuously acts on `weak_dual 𝕜 E`. -/ instance (M) [monoid M] [distrib_mul_action M 𝕜] [smul_comm_class 𝕜 M 𝕜] [topological_space M] [has_continuous_smul M 𝕜] : has_continuous_smul M (weak_dual 𝕜 E) := -⟨continuous_induced_rng $ continuous_fst.smul ((coe_fn_continuous (top_dual_pairing 𝕜 E)).comp - continuous_snd)⟩ +⟨continuous_induced_rng $ continuous_fst.smul ((weak_bilin.coe_fn_continuous + (top_dual_pairing 𝕜 E)).comp continuous_snd)⟩ + +lemma coe_fn_continuous : continuous (λ (x : weak_dual 𝕜 E) y, x y) := +continuous_induced_dom + +lemma eval_continuous (y : E) : continuous (λ x : weak_dual 𝕜 E, x y) := +continuous_pi_iff.mp coe_fn_continuous y + +lemma continuous_of_continuous_eval [topological_space α] {g : α → weak_dual 𝕜 E} + (h : ∀ y, continuous (λ a, (g a) y)) : continuous g := +continuous_induced_rng (continuous_pi_iff.mpr h) + +end weak_dual /-- The weak topology is the topology coarsest topology on `E` such that all functionals `λ x, top_dual_pairing 𝕜 E v x` are continuous. -/ @@ -232,6 +250,6 @@ theorem tendsto_iff_forall_eval_tendsto_top_dual_pairing {l : filter α} {f : α → weak_dual 𝕜 E} {x : weak_dual 𝕜 E} : tendsto f l (𝓝 x) ↔ ∀ y, tendsto (λ i, top_dual_pairing 𝕜 E (f i) y) l (𝓝 (top_dual_pairing 𝕜 E x y)) := -tendsto_iff_forall_eval_tendsto _ continuous_linear_map.coe_injective +weak_bilin.tendsto_iff_forall_eval_tendsto _ continuous_linear_map.coe_injective end weak_star_topology From 163ef61e3d6081bb97ebe9386c713d0d5914e3b6 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Sun, 8 May 2022 18:01:58 +0000 Subject: [PATCH 25/37] feat(topology/algebra/infinite_sum): add `tsum_star` (#13999) These lemmas names are copied from `tsum_neg` and friends. As a result, `star_exp` can be golfed and generalized. --- src/analysis/normed_space/exponential.lean | 18 +++------- .../normed_space/star/exponential.lean | 2 +- src/topology/algebra/infinite_sum.lean | 33 +++++++++++++++++++ 3 files changed, 38 insertions(+), 15 deletions(-) diff --git a/src/analysis/normed_space/exponential.lean b/src/analysis/normed_space/exponential.lean index 1b5d2544146c7..870ab76e23e08 100644 --- a/src/analysis/normed_space/exponential.lean +++ b/src/analysis/normed_space/exponential.lean @@ -125,6 +125,10 @@ begin simp [h] end +lemma star_exp [t2_space 𝔸] [star_ring 𝔸] [has_continuous_star 𝔸] (x : 𝔸) : + star (exp 𝕂 x) = exp 𝕂 (star x) := +by simp_rw [exp_eq_tsum, ←star_pow, ←star_inv_nat_cast_smul, ←tsum_star] + variables (𝕂) lemma commute.exp_right [t2_space 𝔸] {x y : 𝔸} (h : commute x y) : commute x (exp 𝕂 y) := @@ -599,17 +603,3 @@ lemma exp_ℝ_ℂ_eq_exp_ℂ_ℂ : (exp ℝ : ℂ → ℂ) = exp ℂ := exp_eq_exp ℝ ℂ ℂ end scalar_tower - -lemma star_exp {𝕜 A : Type*} [is_R_or_C 𝕜] [normed_ring A] [normed_algebra 𝕜 A] - [star_ring A] [normed_star_group A] [complete_space A] - [star_module 𝕜 A] (a : A) : star (exp 𝕜 a) = exp 𝕜 (star a) := -begin - rw exp_eq_tsum, - have := continuous_linear_map.map_tsum - (starₗᵢ 𝕜 : A ≃ₗᵢ⋆[𝕜] A).to_linear_isometry.to_continuous_linear_map - (exp_series_summable' a), - dsimp at this, - convert this, - funext, - simp only [star_smul, star_pow, one_div, star_inv', star_nat_cast], -end diff --git a/src/analysis/normed_space/star/exponential.lean b/src/analysis/normed_space/star/exponential.lean index 8e56fb4d51985..f7a147543e781 100644 --- a/src/analysis/normed_space/star/exponential.lean +++ b/src/analysis/normed_space/star/exponential.lean @@ -23,7 +23,7 @@ subtypes `self_adjoint A` and `unitary A`. section star variables {A : Type*} -[normed_ring A] [normed_algebra ℂ A] [star_ring A] [cstar_ring A] [complete_space A] +[normed_ring A] [normed_algebra ℂ A] [star_ring A] [has_continuous_star A] [complete_space A] [star_module ℂ A] open complex diff --git a/src/topology/algebra/infinite_sum.lean b/src/topology/algebra/infinite_sum.lean index 04a1ebc14cfec..00071d7fc1c9a 100644 --- a/src/topology/algebra/infinite_sum.lean +++ b/src/topology/algebra/infinite_sum.lean @@ -246,6 +246,26 @@ lemma function.surjective.summable_iff_of_has_sum_iff {α' : Type*} [add_comm_mo summable f ↔ summable g := hes.exists.trans $ exists_congr $ @he +section has_continuous_star +variables [star_add_monoid α] [has_continuous_star α] + +lemma has_sum.star (h : has_sum f a) : has_sum (λ b, star (f b)) (star a) := +by simpa only using h.map (star_add_equiv : α ≃+ α) continuous_star + +lemma summable.star (hf : summable f) : summable (λ b, star (f b)) := +hf.has_sum.star.summable + +lemma summable.of_star (hf : summable (λ b, star (f b))) : summable f := +by simpa only [star_star] using hf.star + +@[simp] lemma summable_star_iff : summable (λ b, star (f b)) ↔ summable f := +⟨summable.of_star, summable.star⟩ + +@[simp] lemma summable_star_iff' : summable (star f) ↔ summable f := +summable_star_iff + +end has_continuous_star + variable [has_continuous_add α] lemma has_sum.add (hf : has_sum f a) (hg : has_sum g b) : has_sum (λb, f b + g b) (a + b) := @@ -474,6 +494,19 @@ end end has_continuous_add +section has_continuous_star +variables [star_add_monoid α] [has_continuous_star α] + +lemma tsum_star : star (∑' b, f b) = ∑' b, star (f b) := +begin + by_cases hf : summable f, + { exact hf.has_sum.star.tsum_eq.symm, }, + { rw [tsum_eq_zero_of_not_summable hf, tsum_eq_zero_of_not_summable (mt summable.of_star hf), + star_zero] }, +end + +end has_continuous_star + section encodable open encodable variable [encodable γ] From 449ba97d90e6bc7023d80771e003c4c950f31f4a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sun, 8 May 2022 19:06:53 +0000 Subject: [PATCH 26/37] refactor(data/nat/log): Golf + improved theorem names (#14019) Other than golfing and moving a few things around, our main changes are: - rename `log_le_log_of_le` to `log_mono_right`, analogous renames elsewhere. - add `lt_pow_iff_log_lt` and a `clog` analog. --- src/data/nat/choose/central.lean | 4 +- src/data/nat/digits.lean | 2 +- src/data/nat/log.lean | 122 +++++++++++++++---------------- src/data/nat/multiplicity.lean | 10 +-- 4 files changed, 65 insertions(+), 73 deletions(-) diff --git a/src/data/nat/choose/central.lean b/src/data/nat/choose/central.lean index 47f7a5d197654..9cd4bf6c27000 100644 --- a/src/data/nat/choose/central.lean +++ b/src/data/nat/choose/central.lean @@ -132,7 +132,7 @@ lemma padic_val_nat_central_binom_of_large_le_one (hp : p.prime) (p_large : 2 * (padic_val_nat p (central_binom n)) ≤ 1 := begin have log_weak_bound : log p (2 * n) ≤ 2, - { calc log p (2 * n) ≤ log p (p ^ 2) : log_le_log_of_le (le_of_lt p_large) + { calc log p (2 * n) ≤ log p (p ^ 2) : log_mono_right (le_of_lt p_large) ... = 2 : log_pow hp.one_lt 2, }, have log_bound : log p (2 * n) ≤ 1, @@ -146,7 +146,7 @@ begin { rw pow_two, exact (nat.le_div_iff_mul_le _ _ (prime.pos hp)).1 v.2.2, }, exact lt_irrefl _ (lt_of_le_of_lt bad p_large), }, - { rw log_eq_zero (or.inl h), + { rw log_of_lt h, exact zero_le 1, }, }, }, exact le_trans (padic_val_nat_central_binom_le hp) log_bound, diff --git a/src/data/nat/digits.lean b/src/data/nat/digits.lean index 624879b666ba5..7cb64d196fd37 100644 --- a/src/data/nat/digits.lean +++ b/src/data/nat/digits.lean @@ -432,7 +432,7 @@ begin { simp [digits_zero_succ', hn] }, { simp, }, { simpa [succ_lt_succ_iff] using hb } }, - simpa [digits_len, hb, hn] using log_le_log_of_le (le_succ _) + simpa [digits_len, hb, hn] using log_mono_right (le_succ _) end lemma le_digits_len_le (b n m : ℕ) (h : n ≤ m) : (digits b n).length ≤ (digits b m).length := diff --git a/src/data/nat/log.lean b/src/data/nat/log.lean index fa01697bdebb5..1d4a5049bf512 100644 --- a/src/data/nat/log.lean +++ b/src/data/nat/log.lean @@ -31,32 +31,28 @@ such that `b^k ≤ n`, so if `b^k = n`, it returns exactly `k`. -/ log (n / b) + 1 else 0 -lemma log_eq_zero {b n : ℕ} (hnb : n < b ∨ b ≤ 1) : log b n = 0 := +private lemma log_eq_zero_aux {b n : ℕ} (hnb : n < b ∨ b ≤ 1) : log b n = 0 := begin rw [or_iff_not_and_not, not_lt, not_le] at hnb, - rw [log, ←ite_not, if_pos hnb], + rw [log, ←ite_not, if_pos hnb] end -lemma log_of_one_lt_of_le {b n : ℕ} (h : 1 < b) (hn : b ≤ n) : log b n = log b (n / b) + 1 := -begin - rw log, - exact if_pos ⟨hn, h⟩, -end +lemma log_of_lt {b n : ℕ} (hb : n < b) : log b n = 0 := +log_eq_zero_aux (or.inl hb) -lemma log_of_lt {b n : ℕ} (hnb : n < b) : log b n = 0 := -by rw [log, if_neg (λ h : b ≤ n ∧ 1 < b, h.1.not_lt hnb)] +lemma log_of_left_le_one {b : ℕ} (hb : b ≤ 1) (n) : log b n = 0 := +log_eq_zero_aux (or.inr hb) -lemma log_of_left_le_one {b n : ℕ} (hb : b ≤ 1) : log b n = 0 := -by rw [log, if_neg (λ h : b ≤ n ∧ 1 < b, h.2.not_le hb)] +lemma log_of_one_lt_of_le {b n : ℕ} (h : 1 < b) (hn : b ≤ n) : log b n = log b (n / b) + 1 := +by { rw log, exact if_pos ⟨hn, h⟩ } lemma log_eq_zero_iff {b n : ℕ} : log b n = 0 ↔ n < b ∨ b ≤ 1 := -begin - refine ⟨λ h_log, _, log_eq_zero⟩, +⟨λ h_log, begin by_contra' h, have := log_of_one_lt_of_le h.2 h.1, rw h_log at this, exact succ_ne_zero _ this.symm -end +end, log_eq_zero_aux⟩ lemma log_eq_one_iff {b n : ℕ} : log b n = 1 ↔ n < b * b ∧ 1 < b ∧ b ≤ n := -- This is best possible: if b = 2, n = 5, then 1 < b and b ≤ n but n > b * b. @@ -78,24 +74,23 @@ begin exact or.inl h, }, end -@[simp] lemma log_zero_left (n : ℕ) : log 0 n = 0 := +@[simp] lemma log_zero_left : ∀ n, log 0 n = 0 := log_of_left_le_one zero_le_one @[simp] lemma log_zero_right (b : ℕ) : log b 0 = 0 := by { rw log, cases b; refl } -@[simp] lemma log_one_left (n : ℕ) : log 1 n = 0 := +@[simp] lemma log_one_left : ∀ n, log 1 n = 0 := log_of_left_le_one le_rfl @[simp] lemma log_one_right (b : ℕ) : log b 1 = 0 := if h : b ≤ 1 then - log_of_left_le_one h + log_of_left_le_one h 1 else log_of_lt (not_le.mp h) /-- `pow b` and `log b` (almost) form a Galois connection. -/ -lemma pow_le_iff_le_log {b : ℕ} (hb : 1 < b) {x y : ℕ} (hy : 0 < y) : - b^x ≤ y ↔ x ≤ log b y := +lemma pow_le_iff_le_log {b : ℕ} (hb : 1 < b) {x y : ℕ} (hy : 0 < y) : b ^ x ≤ y ↔ x ≤ log b y := begin induction y using nat.strong_induction_on with y ih generalizing x, cases x, @@ -109,6 +104,9 @@ begin exact (pow_one b).symm } end +lemma lt_pow_iff_log_lt {b : ℕ} (hb : 1 < b) {x y : ℕ} (hy : 0 < y) : y < b ^ x ↔ log b y < x := +lt_iff_lt_of_le_iff_le (pow_le_iff_le_log hb hy) + lemma log_pow {b : ℕ} (hb : 1 < b) (x : ℕ) : log b (b ^ x) = x := eq_of_forall_le_iff $ λ z, by { rw ←pow_le_iff_le_log hb (pow_pos (zero_lt_one.trans hb) _), @@ -141,7 +139,7 @@ end lemma pow_log_le_self {b : ℕ} (hb : 1 < b) {x : ℕ} (hx : 0 < x) : b ^ log b x ≤ x := (pow_le_iff_le_log hb hx).2 le_rfl -lemma log_le_log_of_le {b n m : ℕ} (h : n ≤ m) : log b n ≤ log b m := +@[mono] lemma log_mono_right {b n m : ℕ} (h : n ≤ m) : log b n ≤ log b m := begin cases le_or_lt b 1 with hb hb, { rw log_of_left_le_one hb, exact zero_le _ }, @@ -151,63 +149,55 @@ begin exact (pow_log_le_self hb hn).trans h } } end -lemma log_le_log_of_left_ge {b c n : ℕ} (hc : 1 < c) (hb : c ≤ b) : log b n ≤ log c n := +@[mono] lemma log_anti_left {b c n : ℕ} (hc : 1 < c) (hb : c ≤ b) : log b n ≤ log c n := begin - cases n, { simp }, - rw ← pow_le_iff_le_log hc (zero_lt_succ n), - calc - c ^ log b n.succ ≤ b ^ log b n.succ : pow_le_pow_of_le_left - (le_of_lt $ zero_lt_one.trans hc) hb _ - ... ≤ n.succ : pow_log_le_self (lt_of_lt_of_le hc hb) - (zero_lt_succ n) + cases n, { rw [log_zero_right, log_zero_right] }, + rw ←pow_le_iff_le_log hc (zero_lt_succ n), + calc c ^ log b n.succ ≤ b ^ log b n.succ : pow_le_pow_of_le_left + (zero_lt_one.trans hc).le hb _ + ... ≤ n.succ : pow_log_le_self (hc.trans_le hb) + (zero_lt_succ n) end -lemma log_monotone {b : ℕ} : monotone (λ n : ℕ, log b n) := -λ x y, log_le_log_of_le +lemma log_monotone {b : ℕ} : monotone (log b) := +λ x y, log_mono_right lemma log_antitone_left {n : ℕ} : antitone_on (λ b, log b n) (set.Ioi 1) := -λ _ hc _ _ hb, log_le_log_of_left_ge (set.mem_Iio.1 hc) hb +λ _ hc _ _ hb, log_anti_left (set.mem_Iio.1 hc) hb @[simp] lemma log_div_mul_self (b n : ℕ) : log b (n / b * b) = log b n := -begin - refine eq_of_forall_le_iff (λ z, _), - split, - { intro h, - exact h.trans (log_monotone (div_mul_le_self _ _)) }, - { intro h, - rcases b with _|_|b, - { simpa using h }, - { simpa using h }, - rcases n.zero_le.eq_or_lt with rfl|hn, - { simpa using h }, - cases le_or_lt b.succ.succ n with hb hb, - { cases z, - { simp }, - have : 0 < b.succ.succ := nat.succ_pos', - rw [←pow_le_iff_le_log, pow_succ'] at h ⊢, - { rwa [(strict_mono_mul_right_of_pos this).le_iff_le, - nat.le_div_iff_mul_le _ _ nat.succ_pos'] }, - all_goals { simp [hn, nat.div_pos hb nat.succ_pos'] } }, - { simpa [div_eq_of_lt, hb, log_eq_zero] using h } } -end +eq_of_forall_le_iff (λ z, ⟨λ h, h.trans (log_monotone (div_mul_le_self _ _)), λ h, begin + rcases b with _|_|b, + { rwa log_zero_left at * }, + { rwa log_one_left at * }, + rcases n.zero_le.eq_or_lt with rfl|hn, + { rwa [nat.zero_div, zero_mul] }, + cases le_or_lt b.succ.succ n with hb hb, + { cases z, + { apply zero_le }, + rw [←pow_le_iff_le_log, pow_succ'] at h ⊢, + { rwa [(strict_mono_mul_right_of_pos nat.succ_pos').le_iff_le, + nat.le_div_iff_mul_le _ _ nat.succ_pos'] }, + all_goals { simp [hn, nat.div_pos hb nat.succ_pos'] } }, + { simpa [div_eq_of_lt, hb, log_of_lt] using h } +end⟩) @[simp] lemma log_div_base (b n : ℕ) : log b (n / b) = log b n - 1 := begin cases lt_or_le n b with h h, - { simp [div_eq_of_lt, h, log_eq_zero] }, + { rw [div_eq_of_lt h, log_of_lt h, log_zero_right] }, rcases n.zero_le.eq_or_lt with rfl|hn, - { simp }, + { rw [nat.zero_div, log_zero_right] }, rcases b with _|_|b, - { simp }, - { simp }, + { rw [log_zero_left, log_zero_left] }, + { rw [log_one_left, log_one_left] }, rw [←succ_inj', ←succ_inj'], simp_rw succ_eq_add_one, rw [nat.sub_add_cancel, ←log_mul_base]; { simp [succ_le_iff, log_pos, h, nat.div_pos] }, end - -private lemma add_pred_div_lt {b n : ℕ} (hb : 1 < b) (hn : 2 ≤ n) : (n + b - 1)/b < n := +private lemma add_pred_div_lt {b n : ℕ} (hb : 1 < b) (hn : 2 ≤ n) : (n + b - 1) / b < n := begin rw [div_lt_iff_lt_mul _ _ (zero_lt_one.trans hb), ←succ_le_iff, ←pred_eq_sub_one, succ_pred_eq_of_pos (add_pos (zero_lt_one.trans hn) (zero_lt_one.trans hb))], @@ -260,8 +250,7 @@ begin end /--`clog b` and `pow b` form a Galois connection. -/ -lemma le_pow_iff_clog_le {b : ℕ} (hb : 1 < b) {x y : ℕ} : - x ≤ b^y ↔ clog b x ≤ y := +lemma le_pow_iff_clog_le {b : ℕ} (hb : 1 < b) {x y : ℕ} : x ≤ b ^ y ↔ clog b x ≤ y := begin induction x using nat.strong_induction_on with x ih generalizing y, cases y, @@ -279,6 +268,9 @@ begin (zero_le _) } end +lemma pow_lt_iff_lt_clog {b : ℕ} (hb : 1 < b) {x y : ℕ} : b ^ y < x ↔ y < clog b x := +lt_iff_lt_of_le_iff_le (le_pow_iff_clog_le hb) + lemma clog_pow (b x : ℕ) (hb : 1 < b) : clog b (b ^ x) = x := eq_of_forall_ge_iff $ λ z, by { rw ←le_pow_iff_clog_le hb, exact (pow_right_strict_mono hb).le_iff_le } @@ -293,7 +285,7 @@ end lemma le_pow_clog {b : ℕ} (hb : 1 < b) (x : ℕ) : x ≤ b ^ clog b x := (le_pow_iff_clog_le hb).2 le_rfl -lemma clog_le_clog_of_le (b : ℕ) {n m : ℕ} (h : n ≤ m) : clog b n ≤ clog b m := +@[mono] lemma clog_mono_right (b : ℕ) {n m : ℕ} (h : n ≤ m) : clog b n ≤ clog b m := begin cases le_or_lt b 1 with hb hb, { rw clog_of_left_le_one hb, exact zero_le _ }, @@ -301,19 +293,19 @@ begin exact h.trans (le_pow_clog hb _) } end -lemma clog_le_clog_of_left_ge {b c n : ℕ} (hc : 1 < c) (hb : c ≤ b) : clog b n ≤ clog c n := +@[mono] lemma clog_anti_left {b c n : ℕ} (hc : 1 < c) (hb : c ≤ b) : clog b n ≤ clog c n := begin rw ← le_pow_iff_clog_le (lt_of_lt_of_le hc hb), calc n ≤ c ^ clog c n : le_pow_clog hc _ - ... ≤ b ^ clog c n : pow_le_pow_of_le_left (le_of_lt $ zero_lt_one.trans hc) hb _ + ... ≤ b ^ clog c n : pow_le_pow_of_le_left (zero_lt_one.trans hc).le hb _ end lemma clog_monotone (b : ℕ) : monotone (clog b) := -λ x y, clog_le_clog_of_le _ +λ x y, clog_mono_right _ lemma clog_antitone_left {n : ℕ} : antitone_on (λ b : ℕ, clog b n) (set.Ioi 1) := -λ _ hc _ _ hb, clog_le_clog_of_left_ge (set.mem_Iio.1 hc) hb +λ _ hc _ _ hb, clog_anti_left (set.mem_Iio.1 hc) hb lemma log_le_clog (b n : ℕ) : log b n ≤ clog b n := begin diff --git a/src/data/nat/multiplicity.lean b/src/data/nat/multiplicity.lean index 1f289fc78e049..916ba94714d98 100644 --- a/src/data/nat/multiplicity.lean +++ b/src/data/nat/multiplicity.lean @@ -97,7 +97,7 @@ lemma multiplicity_factorial {p : ℕ} (hp : p.prime) : calc multiplicity p (n+1)! = multiplicity p n! + multiplicity p (n+1) : by rw [factorial_succ, hp.multiplicity_mul, add_comm] ... = (∑ i in Ico 1 b, n / p ^ i : ℕ) + ((finset.Ico 1 b).filter (λ i, p ^ i ∣ n+1)).card : - by rw [multiplicity_factorial ((log_le_log_of_le $ le_succ _).trans_lt hb), + by rw [multiplicity_factorial ((log_mono_right $ le_succ _).trans_lt hb), ← multiplicity_eq_card_pow_dvd hp.ne_one (succ_pos _) hb] ... = (∑ i in Ico 1 b, (n / p ^ i + if p^i ∣ n+1 then 1 else 0) : ℕ) : by { rw [sum_add_distrib, sum_boole], simp } @@ -178,8 +178,8 @@ have h₁ : multiplicity p (choose n k) + multiplicity p (k! * (n - k)!) = begin rw [← hp.multiplicity_mul, ← mul_assoc, choose_mul_factorial_mul_factorial hkn, hp.multiplicity_factorial hnb, hp.multiplicity_mul, - hp.multiplicity_factorial ((log_le_log_of_le hkn).trans_lt hnb), - hp.multiplicity_factorial (lt_of_le_of_lt (log_le_log_of_le tsub_le_self) hnb), + hp.multiplicity_factorial ((log_mono_right hkn).trans_lt hnb), + hp.multiplicity_factorial (lt_of_le_of_lt (log_mono_right tsub_le_self) hnb), multiplicity_choose_aux hp hkn], simp [add_comm], end, @@ -198,7 +198,7 @@ else if hn0 : n = 0 then by cases k; simp [hn0, *] at * else begin rw [multiplicity_choose hp (le_of_not_gt hkn) (lt_succ_self _), multiplicity_eq_card_pow_dvd (ne_of_gt hp.one_lt) (nat.pos_of_ne_zero hk0) - (lt_succ_of_le (log_le_log_of_le (le_of_not_gt hkn))), + (lt_succ_of_le (log_mono_right (le_of_not_gt hkn))), multiplicity_eq_card_pow_dvd (ne_of_gt hp.one_lt) (nat.pos_of_ne_zero hn0) (lt_succ_self _), ← nat.cast_add, enat.coe_le_coe], calc ((Ico 1 (log p n).succ).filter (λ i, p ^ i ∣ n)).card @@ -226,7 +226,7 @@ le_antisymm begin rw [multiplicity_choose hp hkn (lt_succ_self _), multiplicity_eq_card_pow_dvd (ne_of_gt hp.one_lt) hk0 - (lt_succ_of_le (log_le_log_of_le hkn)), + (lt_succ_of_le (log_mono_right hkn)), ← nat.cast_add, enat.coe_le_coe, log_pow hp.one_lt, ← card_disjoint_union hdisj, filter_union_right], have filter_le_Ico := (Ico 1 n.succ).card_filter_le _, From bf6e13bf9e257d52d6216afddeffee7e7280ac08 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 9 May 2022 00:34:40 +0000 Subject: [PATCH 27/37] refactor(algebra/{group,group_with_zero/basic): Generalize lemmas to division monoids (#14000) Generalize `group` and `group_with_zero` lemmas to `division_monoid`. We do not actually delete the original lemmas but make them one-liners from the new ones. The next PR will then delete the old lemmas and perform the renames in all files. Lemmas are renamed because * one of the `group` or `group_with_zero` name has to go * the new API should have a consistent naming convention Pre-emptive lemma renames --- .../computation/terminates_iff_rat.lean | 3 +- src/algebra/group/basic.lean | 201 ++++++++++++------ src/algebra/group/defs.lean | 3 + src/algebra/group_with_zero/basic.lean | 106 ++++----- .../trigonometric/complex.lean | 4 +- 5 files changed, 176 insertions(+), 141 deletions(-) diff --git a/src/algebra/continued_fractions/computation/terminates_iff_rat.lean b/src/algebra/continued_fractions/computation/terminates_iff_rat.lean index 629a929e26383..92de5b24686e4 100644 --- a/src/algebra/continued_fractions/computation/terminates_iff_rat.lean +++ b/src/algebra/continued_fractions/computation/terminates_iff_rat.lean @@ -188,8 +188,7 @@ begin have : (fr : K)⁻¹ = ((fr⁻¹ : ℚ) : K), by norm_cast, have coe_of_fr := (coe_of_rat_eq this), simp [int_fract_pair.stream, IH.symm, v_eq_q, stream_q_nth_eq, fr_ne_zero], - unfold_coes, - simpa [coe_of_fr] } } } + exact congr_arg some coe_of_fr } } } end lemma coe_stream_rat_eq : diff --git a/src/algebra/group/basic.lean b/src/algebra/group/basic.lean index 67a04a42de71a..2b802531b8dfb 100644 --- a/src/algebra/group/basic.lean +++ b/src/algebra/group/basic.lean @@ -16,10 +16,13 @@ one-liners from the corresponding axioms. For the definitions of semigroups, mon `algebra/group/defs.lean`. -/ +open function + universe u +variables {α G : Type*} section associative -variables {α : Type u} (f : α → α → α) [is_associative α f] (x y : α) +variables (f : α → α → α) [is_associative α f] (x y : α) /-- Composing two associative operations of `f : α → α → α` on the left @@ -38,7 +41,6 @@ by { ext z, rw [function.comp_apply, @is_associative.assoc _ f] } end associative section semigroup -variables {α : Type*} /-- Composing two multiplications on the left by `y` then `x` @@ -90,7 +92,7 @@ lemma mul_one_eq_id : (* (1 : M)) = id := funext mul_one end mul_one_class section comm_semigroup -variables {G : Type u} [comm_semigroup G] +variables [comm_semigroup G] @[no_rsimp, to_additive] lemma mul_left_comm : ∀ a b c : G, a * (b * c) = b * (a * c) := @@ -160,7 +162,7 @@ eq_comm.trans mul_left_eq_self end right_cancel_monoid section has_involutive_inv -variables {G : Type u} [has_involutive_inv G] {a b : G} +variables [has_involutive_inv G] {a b : G} @[simp, to_additive] lemma inv_involutive : function.involutive (has_inv.inv : G → G) := inv_inv @@ -187,11 +189,15 @@ theorem eq_inv_iff_eq_inv : a = b⁻¹ ↔ b = a⁻¹ := theorem inv_eq_iff_inv_eq : a⁻¹ = b ↔ b⁻¹ = a := eq_comm.trans $ eq_inv_iff_eq_inv.trans eq_comm +variables (G) + +@[to_additive] lemma left_inverse_inv : left_inverse (λ a : G, a⁻¹) (λ a, a⁻¹) := inv_inv +@[to_additive] lemma right_inverse_inv : left_inverse (λ a : G, a⁻¹) (λ a, a⁻¹) := inv_inv + end has_involutive_inv section div_inv_monoid - -variables {G : Type u} [div_inv_monoid G] +variables [div_inv_monoid G] {a b c : G} @[to_additive, field_simps] -- The attributes are out of order on purpose lemma inv_eq_one_div (x : G) : @@ -218,28 +224,106 @@ lemma mul_div_assoc' (a b c : G) : a * (b / c) = (a * b) / c := lemma mul_div (a b c : G) : a * (b / c) = a * b / c := by simp only [mul_assoc, div_eq_mul_inv] +@[to_additive] lemma div_eq_mul_one_div (a b : G) : a / b = a * (1 / b) := +by rw [div_eq_mul_inv, one_div] + end div_inv_monoid -section group -variables {G : Type u} [group G] {a b c d : G} +namespace division_monoid +variables [division_monoid α] {a b c : α} -@[simp, to_additive neg_zero] -lemma one_inv : 1⁻¹ = (1 : G) := -inv_eq_of_mul_eq_one (one_mul 1) +local attribute [simp] mul_assoc div_eq_mul_inv -@[to_additive] -theorem left_inverse_inv (G) [has_involutive_inv G] : - function.left_inverse (λ a : G, a⁻¹) (λ a, a⁻¹) := -inv_inv +@[to_additive] lemma inv_eq_of_mul_eq_one_left (h : a * b = 1) : b⁻¹ = a := +by rw [←inv_eq_of_mul_eq_one_right h, inv_inv] -@[simp, to_additive] -theorem inv_eq_one : a⁻¹ = 1 ↔ a = 1 := inv_injective.eq_iff' one_inv +@[to_additive] lemma eq_inv_of_mul_eq_one_left (h : a * b = 1) : a = b⁻¹ := +(inv_eq_of_mul_eq_one_left h).symm -@[simp, to_additive] -theorem one_eq_inv : 1 = a⁻¹ ↔ a = 1 := eq_comm.trans inv_eq_one +@[to_additive] lemma eq_inv_of_mul_eq_one_right (h : a * b = 1) : b = a⁻¹ := +(inv_eq_of_mul_eq_one_right h).symm -@[to_additive] -theorem inv_ne_one : a⁻¹ ≠ 1 ↔ a ≠ 1 := not_congr inv_eq_one +@[to_additive] lemma eq_one_div_of_mul_eq_one_left (h : b * a = 1) : b = 1 / a := +by rw [eq_inv_of_mul_eq_one_left h, one_div] + +@[to_additive] lemma eq_one_div_of_mul_eq_one_right (h : a * b = 1) : b = 1 / a := +by rw [eq_inv_of_mul_eq_one_right h, one_div] + +@[to_additive] lemma eq_of_div_eq_one (h : a / b = 1) : a = b := +inv_injective $ inv_eq_of_mul_eq_one_right $ by rwa ←div_eq_mul_inv + +@[to_additive] lemma div_ne_one_of_ne : a ≠ b → a / b ≠ 1 := mt eq_of_div_eq_one + +variables (a b c) + +@[to_additive] lemma one_div_mul_one_div_rev : (1 / a) * (1 / b) = 1 / (b * a) := by simp +@[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 + +variables {a b c} + +@[simp, to_additive] lemma inv_eq_one : a⁻¹ = 1 ↔ a = 1 := inv_injective.eq_iff' inv_one +@[simp, to_additive] lemma one_eq_inv : 1 = a⁻¹ ↔ a = 1 := eq_comm.trans inv_eq_one +@[to_additive] lemma inv_ne_one : a⁻¹ ≠ 1 ↔ a ≠ 1 := inv_eq_one.not + +@[to_additive] lemma eq_of_one_div_eq_one_div (h : 1 / a = 1 / b) : a = b := +by rw [←one_div_one_div a, h, one_div_one_div] + +variables (a b c) + + -- The attributes are out of order on purpose +@[to_additive, field_simps] lemma div_div_eq_mul_div : a / (b / c) = a * c / b := by simp +@[simp, to_additive] lemma div_inv_eq_mul : a / b⁻¹ = a * b := by simp +@[to_additive] lemma div_mul_eq_div_div_swap : a / (b * c) = a / c / b := +by simp only [mul_assoc, mul_inv_rev, div_eq_mul_inv] + +end division_monoid + +namespace division_comm_monoid +variables [division_comm_monoid α] (a b c d : α) + +local attribute [simp] mul_assoc mul_comm mul_left_comm div_eq_mul_inv + +@[to_additive neg_add] lemma mul_inv : (a * b)⁻¹ = a⁻¹ * b⁻¹ := by simp +@[to_additive] lemma div_eq_inv_mul : a / b = b⁻¹ * a := by simp +@[to_additive] lemma inv_mul_eq_div : a⁻¹ * b = b / a := by simp +@[to_additive] lemma inv_mul' : (a * b)⁻¹ = a⁻¹ / b := by simp +@[to_additive] lemma inv_div_inv : (a⁻¹ / b⁻¹) = b / a := by simp +@[to_additive] lemma inv_inv_div_inv : (a⁻¹ / b⁻¹)⁻¹ = a / b := by simp +@[to_additive] lemma one_div_mul_one_div : (1 / a) * (1 / b) = 1 / (a * b) := by simp + +@[to_additive] lemma div_right_comm : a / b / c = a / c / b := by simp +@[to_additive] lemma div_div : a / b / c = a / (b * c) := by simp +@[to_additive] lemma div_mul : a / b * c = a / (b / c) := by simp +@[to_additive] lemma mul_div_left_comm : a * (b / c) = b * (a / c) := by simp +@[to_additive] lemma mul_div_right_comm : a * b / c = a / c * b := by simp +@[to_additive] lemma div_mul_eq_div_div : a / (b * c) = a / b / c := by simp +@[to_additive, field_simps] lemma div_mul_eq_mul_div : a / b * c = a * c / b := by simp +@[to_additive] lemma mul_comm_div : a / b * c = a * (c / b) := by simp +@[to_additive] lemma div_mul_comm : a / b * c = c / b * a := by simp +@[to_additive] lemma div_mul_eq_div_mul_one_div : a / (b * c) = (a / b) * (1 / c) := by simp + +@[to_additive] lemma div_div_div_eq : a / b / (c / d) = a * d / (b * c) := by simp +@[to_additive] lemma div_div_div_comm : a / b / (c / d) = a / c / (b / d) := by simp +@[to_additive] lemma div_mul_div_comm : a / b * (c / d) = a * c / (b * d) := by simp +@[to_additive] lemma mul_div_mul_comm : a * b / (c * d) = a / c * (b / d) := by simp + +end division_comm_monoid + +section group +variables [group G] {a b c d : G} + +@[to_additive neg_zero] lemma one_inv : 1⁻¹ = (1 : G) := division_monoid.inv_one + +@[to_additive] theorem inv_eq_one : a⁻¹ = 1 ↔ a = 1 := division_monoid.inv_eq_one +@[to_additive] theorem one_eq_inv : 1 = a⁻¹ ↔ a = 1 := division_monoid.one_eq_inv +@[to_additive] theorem inv_ne_one : a⁻¹ ≠ 1 ↔ a ≠ 1 := division_monoid.inv_ne_one @[simp, to_additive] theorem div_eq_inv_self : a / b = b⁻¹ ↔ a = 1 := by rw [div_eq_mul_inv, mul_left_eq_self] @@ -253,9 +337,7 @@ theorem mul_right_surjective (a : G) : function.surjective (λ x, x * a) := λ x, ⟨x * a⁻¹, inv_mul_cancel_right x a⟩ @[to_additive] -lemma eq_inv_of_mul_eq_one (h : a * b = 1) : a = b⁻¹ := -have a⁻¹ = b, from inv_eq_of_mul_eq_one h, -by simp [this.symm] +lemma eq_inv_of_mul_eq_one : a * b = 1 → a = b⁻¹ := division_monoid.eq_inv_of_mul_eq_one_left @[to_additive] lemma eq_mul_inv_of_mul_eq (h : a * c = b) : a = b * c⁻¹ := @@ -337,9 +419,8 @@ by simpa only [div_eq_mul_inv] using λ a a' h, mul_left_injective (b⁻¹) h lemma div_right_injective : function.injective (λ a, b / a) := by simpa only [div_eq_mul_inv] using λ a a' h, inv_injective (mul_right_injective b h) -@[simp, to_additive neg_sub] -lemma inv_div' (a b : G) : (a / b)⁻¹ = b / a := -by rw [div_eq_mul_inv, div_eq_mul_inv, mul_inv_rev, inv_inv] +@[to_additive neg_sub] +lemma inv_div' (a b : G) : (a / b)⁻¹ = b / a := division_monoid.inv_div _ _ @[simp, to_additive sub_add_cancel] lemma div_mul_cancel' (a b : G) : a / b * b = a := @@ -354,21 +435,16 @@ lemma mul_div_cancel'' (a b : G) : a * b / b = a := by rw [div_eq_mul_inv, mul_inv_cancel_right a b] @[to_additive eq_of_sub_eq_zero] -lemma eq_of_div_eq_one' (h : a / b = 1) : a = b := -calc a = a / b * b : (div_mul_cancel' a b).symm - ... = b : by rw [h, one_mul] +lemma eq_of_div_eq_one' : a / b = 1 → a = b := division_monoid.eq_of_div_eq_one -@[to_additive] -lemma div_ne_one_of_ne (h : a ≠ b) : a / b ≠ 1 := -mt eq_of_div_eq_one' h +@[to_additive] lemma div_ne_one_of_ne : a ≠ b → a / b ≠ 1 := division_monoid.div_ne_one_of_ne -@[simp, to_additive] -lemma div_inv_eq_mul (a b : G) : a / (b⁻¹) = a * b := -by rw [div_eq_mul_inv, inv_inv] +@[to_additive] +lemma div_inv_eq_mul (a b : G) : a / (b⁻¹) = a * b := division_monoid.div_inv_eq_mul _ _ @[to_additive] lemma div_mul_eq_div_div_swap (a b c : G) : a / (b * c) = a / c / b := -by simp only [mul_assoc, mul_inv_rev , div_eq_mul_inv] +division_monoid.div_mul_eq_div_div_swap _ _ _ @[simp, to_additive] lemma mul_div_mul_right_eq_div (a b c : G) : (a * c) / (b * c) = a / b := @@ -407,8 +483,7 @@ lemma div_div_div_cancel_right' (a b c : G) : (a / c) / (b / c) = a / b := by rw [← inv_div' c b, div_inv_eq_mul, div_mul_div_cancel'] @[to_additive] -theorem div_div_assoc_swap : a / (b / c) = a * c / b := -by simp only [mul_assoc, mul_inv_rev, inv_inv, div_eq_mul_inv] +theorem div_div_assoc_swap : a / (b / c) = a * c / b := division_monoid.div_div_eq_mul_div _ _ _ @[to_additive] theorem div_eq_one : a / b = 1 ↔ a = b := @@ -427,9 +502,7 @@ by rw [div_eq_mul_inv, mul_right_eq_self, inv_eq_one] -- The unprimed version is used by `group_with_zero`. This is the preferred choice. -- See https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/.60div_one'.60 -@[simp, to_additive sub_zero] -lemma div_one' (a : G) : a / 1 = a := -div_eq_self.2 rfl +@[to_additive sub_zero] lemma div_one' (a : G) : a / 1 = a := division_monoid.div_one _ @[to_additive eq_sub_iff_add_eq] theorem eq_div_iff_mul_eq' : a = b / c ↔ a * c = b := @@ -475,13 +548,12 @@ end end group section comm_group -variables {G : Type u} [comm_group G] {a b c d : G} +variables [comm_group G] {a b c d : G} local attribute [simp] mul_assoc mul_comm mul_left_comm div_eq_mul_inv @[to_additive neg_add] -lemma mul_inv (a b : G) : (a * b)⁻¹ = a⁻¹ * b⁻¹ := -by rw [mul_inv_rev, mul_comm] +lemma mul_inv (a b : G) : (a * b)⁻¹ = a⁻¹ * b⁻¹ := division_comm_monoid.mul_inv _ _ @[to_additive] lemma div_eq_of_eq_mul' {a b c : G} (h : a = b * c) : a / b = c := @@ -489,34 +561,32 @@ by rw [h, div_eq_mul_inv, mul_comm, inv_mul_cancel_left] @[to_additive] lemma mul_div_left_comm {x y z : G} : x * (y / z) = y * (x / z) := -by simp_rw [div_eq_mul_inv, mul_left_comm] +division_comm_monoid.mul_div_left_comm _ _ _ @[to_additive] lemma div_mul_div_comm (a b c d : G) : a / b * (c / d) = a * c / (b * d) := -by simp +division_comm_monoid.div_mul_div_comm _ _ _ _ @[to_additive] -lemma div_div_div_comm (a b c d : G) : (a / b) / (c / d) = (a / c) / (b / d) := by simp +lemma div_div_div_comm (a b c d : G) : (a / b) / (c / d) = (a / c) / (b / d) := +division_comm_monoid.div_div_div_comm _ _ _ _ @[to_additive] lemma div_mul_eq_div_div (a b c : G) : a / (b * c) = a / b / c := -by simp +division_comm_monoid.div_mul_eq_div_div _ _ _ @[to_additive] -lemma inv_mul_eq_div (a b : G) : a⁻¹ * b = b / a := -by simp +lemma inv_mul_eq_div (a b : G) : a⁻¹ * b = b / a := division_comm_monoid.inv_mul_eq_div _ _ @[to_additive sub_add_eq_add_sub] lemma div_mul_eq_mul_div' (a b c : G) : a / b * c = a * c / b := -by simp +division_comm_monoid.div_mul_eq_mul_div _ _ _ @[to_additive] -lemma div_div (a b c : G) : a / b / c = a / (b * c) := -by simp +lemma div_div (a b c : G) : a / b / c = a / (b * c) := division_comm_monoid.div_div _ _ _ @[to_additive] -lemma div_mul (a b c : G) : a / b * c = a / (b / c) := -by simp +lemma div_mul (a b c : G) : a / b * c = a / (b / c) := division_comm_monoid.div_mul _ _ _ @[simp, to_additive] lemma mul_div_mul_left_eq_div (a b c : G) : (c * a) / (c * b) = a / b := @@ -540,33 +610,28 @@ by simpa using mul_inv_cancel_left a b @[to_additive add_sub_comm] lemma mul_div_comm' (a b c d : G) : a * b / (c * d) = (a / c) * (b / d) := -by simp +division_comm_monoid.mul_div_mul_comm _ _ _ _ @[to_additive] -lemma div_eq_div_mul_div (a b c : G) : a / b = c / b * (a / c) := -begin simp, rw [mul_left_comm c], simp end +lemma div_eq_div_mul_div (a b c : G) : a / b = c / b * (a / c) := by simp [mul_left_comm c] @[to_additive] -lemma inv_inv_div_inv (a b : G) : (a⁻¹ / b⁻¹)⁻¹ = a / b := -by simp +lemma inv_inv_div_inv (a b : G) : (a⁻¹ / b⁻¹)⁻¹ = a / b := division_comm_monoid.inv_inv_div_inv _ _ @[simp, to_additive] lemma div_div_cancel (a b : G) : a / (a / b) = b := div_div_self' a b @[to_additive sub_eq_neg_add] -lemma div_eq_inv_mul' (a b : G) : a / b = b⁻¹ * a := -by rw [div_eq_mul_inv, mul_comm _ _] +lemma div_eq_inv_mul' (a b : G) : a / b = b⁻¹ * a := division_comm_monoid.div_eq_inv_mul _ _ @[simp, to_additive] lemma div_div_cancel_left (a b : G) : a / b / a = b⁻¹ := by simp @[to_additive] -theorem inv_mul' (a b : G) : (a * b)⁻¹ = a⁻¹ / b := -by rw [div_eq_mul_inv, mul_inv a b] +theorem inv_mul' (a b : G) : (a * b)⁻¹ = a⁻¹ / b := division_comm_monoid.inv_mul' _ _ @[simp, to_additive] -lemma inv_div_inv (a b : G) : a⁻¹ / b⁻¹ = b / a := -by simp [div_eq_inv_mul', mul_comm] +lemma inv_div_inv (a b : G) : a⁻¹ / b⁻¹ = b / a := division_comm_monoid.inv_div_inv _ _ @[to_additive eq_sub_iff_add_eq'] lemma eq_div_iff_mul_eq'' : a = b / c ↔ c * a = b := @@ -597,7 +662,7 @@ by rw [← div_eq_mul_inv, mul_div_cancel'_right a b] @[to_additive sub_right_comm] lemma div_right_comm' (a b c : G) : a / b / c = a / c / b := -by { repeat { rw div_eq_mul_inv }, exact mul_right_comm _ _ _ } +division_comm_monoid.div_right_comm _ _ _ @[simp, to_additive] lemma mul_mul_div_cancel (a b c : G) : (a * c) * (b / c) = a * b := diff --git a/src/algebra/group/defs.lean b/src/algebra/group/defs.lean index 419130e2da619..380882efa4c5a 100644 --- a/src/algebra/group/defs.lean +++ b/src/algebra/group/defs.lean @@ -685,6 +685,9 @@ variables [division_monoid G] {a b : G} @[simp, to_additive neg_add_rev] lemma mul_inv_rev (a b : G) : (a * b)⁻¹ = b⁻¹ * a⁻¹ := division_monoid.mul_inv_rev _ _ +@[to_additive] +lemma inv_eq_of_mul_eq_one_right : a * b = 1 → a⁻¹ = b := division_monoid.inv_eq_of_mul _ _ + @[simp, to_additive] lemma inv_eq_of_mul_eq_one : a * b = 1 → a⁻¹ = b := division_monoid.inv_eq_of_mul _ _ diff --git a/src/algebra/group_with_zero/basic.lean b/src/algebra/group_with_zero/basic.lean index ac701b3712324..649717c0769a9 100644 --- a/src/algebra/group_with_zero/basic.lean +++ b/src/algebra/group_with_zero/basic.lean @@ -571,10 +571,6 @@ calc (a * b⁻¹) * b = a * (b⁻¹ * b) : mul_assoc _ _ _ calc a⁻¹ * (a * b) = (a⁻¹ * a) * b : (mul_assoc _ _ _).symm ... = b : by simp [h] -@[simp] lemma inv_one : 1⁻¹ = (1:G₀) := -calc 1⁻¹ = 1 * 1⁻¹ : by rw [one_mul] - ... = (1:G₀) : by simp - private lemma inv_eq_of_mul (h : a * b = 1) : a⁻¹ = b := by rw [← inv_mul_cancel_left₀ (left_ne_zero_of_mul_eq_one h) b, h, mul_one] @@ -595,6 +591,8 @@ instance group_with_zero.to_division_monoid : division_monoid G₀ := inv_eq_of_mul := λ a b, inv_eq_of_mul, ..‹group_with_zero G₀› } +@[simp] lemma inv_one : 1⁻¹ = (1:G₀) := division_monoid.inv_one + /-- Multiplying `a` by itself and then by its inverse results in `a` (whether or not `a` is zero). -/ @[simp] lemma mul_self_mul_inv (a : G₀) : a * a * a⁻¹ = a := @@ -632,16 +630,10 @@ by rw [div_eq_mul_inv, mul_self_mul_inv a] @[simp] lemma div_self_mul_self (a : G₀) : a / a * a = a := by rw [div_eq_mul_inv, mul_inv_mul_self a] -lemma eq_inv_of_mul_right_eq_one (h : a * b = 1) : - b = a⁻¹ := -by rw [← inv_mul_cancel_left₀ (left_ne_zero_of_mul_eq_one h) b, h, mul_one] - -lemma eq_inv_of_mul_left_eq_one (h : a * b = 1) : - a = b⁻¹ := -by rw [← mul_inv_cancel_right₀ (right_ne_zero_of_mul_eq_one h) a, h, one_mul] +lemma eq_inv_of_mul_right_eq_one : a * b = 1 → b = a⁻¹ := division_monoid.eq_inv_of_mul_eq_one_right +lemma eq_inv_of_mul_left_eq_one : a * b = 1 → a = b⁻¹ := division_monoid.eq_inv_of_mul_eq_one_left -@[simp] lemma inv_eq_one₀ : g⁻¹ = 1 ↔ g = 1 := -by rw [inv_eq_iff_inv_eq, inv_one, eq_comm] +lemma inv_eq_one₀ : g⁻¹ = 1 ↔ g = 1 := division_monoid.inv_eq_one lemma eq_mul_inv_iff_mul_eq₀ (hc : c ≠ 0) : a = b * c⁻¹ ↔ a * c = b := by split; rintro rfl; [rw inv_mul_cancel_right₀ hc, rw mul_inv_cancel_right₀ hc] @@ -761,20 +753,12 @@ instance group_with_zero.cancel_monoid_with_zero : cancel_monoid_with_zero G₀ units.mk0 x (mul_ne_zero_iff.mp hxy).1 * units.mk0 y (mul_ne_zero_iff.mp hxy).2 := by { ext, refl } -lemma mul_inv_rev₀ (x y : G₀) : (x * y)⁻¹ = y⁻¹ * x⁻¹ := -begin - by_cases hx : x = 0, { simp [hx] }, - by_cases hy : y = 0, { simp [hy] }, - symmetry, - apply eq_inv_of_mul_left_eq_one, - simp [mul_assoc, hx, hy] -end +lemma mul_inv_rev₀ (x y : G₀) : (x * y)⁻¹ = y⁻¹ * x⁻¹ := mul_inv_rev _ _ @[simp] lemma div_self {a : G₀} (h : a ≠ 0) : a / a = 1 := by rw [div_eq_mul_inv, mul_inv_cancel h] -@[simp] lemma div_one (a : G₀) : a / 1 = a := -by simp [div_eq_mul_inv a 1] +lemma div_one (a : G₀) : a / 1 = a := division_monoid.div_one _ @[simp] lemma zero_div (a : G₀) : 0 / a = 0 := by rw [div_eq_mul_inv, zero_mul] @@ -800,35 +784,29 @@ local attribute [simp] div_eq_mul_inv mul_comm mul_assoc mul_left_comm calc a / (a * a) = a⁻¹⁻¹ * a⁻¹ * a⁻¹ : by simp [mul_inv_rev₀] ... = a⁻¹ : inv_mul_mul_self _ -lemma div_eq_mul_one_div (a b : G₀) : a / b = a * (1 / b) := -by simp - lemma mul_one_div_cancel {a : G₀} (h : a ≠ 0) : a * (1 / a) = 1 := by simp [h] lemma one_div_mul_cancel {a : G₀} (h : a ≠ 0) : (1 / a) * a = 1 := by simp [h] -lemma one_div_one : 1 / 1 = (1:G₀) := -div_self (ne.symm zero_ne_one) +lemma one_div_one : 1 / 1 = (1:G₀) := division_monoid.one_div_one lemma one_div_ne_zero {a : G₀} (h : a ≠ 0) : 1 / a ≠ 0 := by simpa only [one_div] using inv_ne_zero h -lemma eq_one_div_of_mul_eq_one {a b : G₀} (h : a * b = 1) : b = 1 / a := -by simpa only [one_div] using eq_inv_of_mul_right_eq_one h +lemma eq_one_div_of_mul_eq_one {a b : G₀} : a * b = 1 → b = 1 / a := +division_monoid.eq_one_div_of_mul_eq_one_right -lemma eq_one_div_of_mul_eq_one_left {a b : G₀} (h : b * a = 1) : b = 1 / a := -by simpa only [one_div] using eq_inv_of_mul_left_eq_one h +lemma eq_one_div_of_mul_eq_one_left {a b : G₀} : b * a = 1 → b = 1 / a := +division_monoid.eq_one_div_of_mul_eq_one_left -@[simp] lemma one_div_div (a b : G₀) : 1 / (a / b) = b / a := -by rw [one_div, div_eq_mul_inv, mul_inv_rev₀, inv_inv, div_eq_mul_inv] +lemma one_div_div (a b : G₀) : 1 / (a / b) = b / a := division_monoid.one_div_div _ _ -lemma one_div_one_div (a : G₀) : 1 / (1 / a) = a := -by simp +lemma one_div_one_div (a : G₀) : 1 / (1 / a) = a := division_monoid.one_div_one_div _ -lemma eq_of_one_div_eq_one_div {a b : G₀} (h : 1 / a = 1 / b) : a = b := -by rw [← one_div_one_div a, h, one_div_one_div] +lemma eq_of_one_div_eq_one_div {a b : G₀} : 1 / a = 1 / b → a = b := +division_monoid.eq_of_one_div_eq_one_div variables {a b c : G₀} @@ -839,7 +817,7 @@ by rw [inv_eq_iff_inv_eq, inv_zero, eq_comm] eq_comm.trans $ inv_eq_zero.trans eq_comm lemma one_div_mul_one_div_rev (a b : G₀) : (1 / a) * (1 / b) = 1 / (b * a) := -by simp only [div_eq_mul_inv, one_mul, mul_inv_rev₀] +division_monoid.one_div_mul_one_div_rev _ _ theorem divp_eq_div (a : G₀) (u : G₀ˣ) : a /ₚ u = a / u := by simpa only [div_eq_mul_inv] using congr_arg ((*) a) u.coe_inv' @@ -848,11 +826,9 @@ by simpa only [div_eq_mul_inv] using congr_arg ((*) a) u.coe_inv' a /ₚ units.mk0 b hb = a / b := divp_eq_div _ _ -lemma inv_div : (a / b)⁻¹ = b / a := -by rw [div_eq_mul_inv, mul_inv_rev₀, div_eq_mul_inv, inv_inv] +lemma inv_div : (a / b)⁻¹ = b / a := division_monoid.inv_div _ _ -lemma inv_div_left : a⁻¹ / b = (b * a)⁻¹ := -by rw [mul_inv_rev₀, div_eq_mul_inv] +lemma inv_div_left : a⁻¹ / b = (b * a)⁻¹ := division_monoid.inv_div_left _ _ 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) } @@ -879,13 +855,7 @@ lemma div_eq_of_eq_mul {x : G₀} (hx : x ≠ 0) {y z : G₀} (h : y = z * x) : lemma eq_div_of_mul_eq {x : G₀} (hx : x ≠ 0) {y z : G₀} (h : z * x = y) : z = y / x := eq.symm $ div_eq_of_eq_mul hx h.symm -lemma eq_of_div_eq_one (h : a / b = 1) : a = b := -begin - by_cases hb : b = 0, - { rw [hb, div_zero] at h, - exact eq_of_zero_eq_one h a b }, - { rwa [div_eq_iff_mul_eq hb, one_mul, eq_comm] at h } -end +lemma eq_of_div_eq_one : a / b = 1 → a = b := division_monoid.eq_of_div_eq_one lemma div_eq_one_iff_eq (hb : b ≠ 0) : a / b = 1 ↔ a = b := ⟨eq_of_div_eq_one, λ h, h.symm ▸ div_self hb⟩ @@ -912,7 +882,7 @@ funext ring.inverse_eq_inv @[field_simps] lemma div_div_eq_mul_div (a b c : G₀) : a / (b / c) = (a * c) / b := -by rw [div_eq_mul_one_div, one_div_div, ← mul_div_assoc] +division_monoid.div_div_eq_mul_div _ _ _ /-- Dividing `a` by the result of dividing `a` by itself results in `a` (whether or not `a` is zero). -/ @@ -977,11 +947,10 @@ 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 mul_inv₀ : (a * b)⁻¹ = a⁻¹ * b⁻¹ := -by rw [mul_inv_rev₀, mul_comm] +lemma mul_inv₀ : (a * b)⁻¹ = a⁻¹ * b⁻¹ := division_comm_monoid.mul_inv _ _ lemma one_div_mul_one_div (a b : G₀) : (1 / a) * (1 / b) = 1 / (a * b) := -by rw [one_div_mul_one_div_rev, mul_comm b] +division_comm_monoid.one_div_mul_one_div _ _ lemma div_mul_right {a : G₀} (b : G₀) (ha : a ≠ 0) : a / (a * b) = 1 / b := by rw [mul_comm, div_mul_left ha] @@ -1002,21 +971,21 @@ local attribute [simp] mul_assoc mul_comm mul_left_comm lemma div_mul_div_comm₀ (a b c d : G₀) : (a / b) * (c / d) = (a * c) / (b * d) := -by simp [div_eq_mul_inv, mul_inv₀] +division_comm_monoid.div_mul_div_comm _ _ _ _ lemma div_div_div_comm₀ (a b c d : G₀) : (a / b) / (c / d) = (a / c) / (b / d) := -by simp_rw [div_eq_mul_inv, mul_inv₀, inv_inv, mul_mul_mul_comm] +division_comm_monoid.div_div_div_comm _ _ _ _ lemma mul_div_mul_left (a b : G₀) {c : G₀} (hc : c ≠ 0) : (c * a) / (c * b) = a / b := by rw [mul_comm c, mul_comm c, mul_div_mul_right _ _ hc] @[field_simps] lemma div_mul_eq_mul_div (a b c : G₀) : (b / c) * a = (b * a) / c := -by simp [div_eq_mul_inv] +division_comm_monoid.div_mul_eq_mul_div _ _ _ lemma div_mul_eq_mul_div_comm (a b c : G₀) : (b / c) * a = b * (a / c) := -by rw [div_mul_eq_mul_div, ← one_mul c, ← div_mul_div_comm₀, div_one, one_mul] +division_comm_monoid.mul_comm_div _ _ _ 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 := @@ -1025,15 +994,15 @@ by rw [← mul_one (a*d), mul_assoc, mul_comm d, ← mul_assoc, ← div_self hb, @[field_simps] lemma div_div_eq_div_mul (a b c : G₀) : (a / b) / c = a / (b * c) := -by rw [div_eq_mul_one_div, div_mul_div_comm₀, mul_one] +division_comm_monoid.div_div _ _ _ lemma div_div_div_div_eq (a : G₀) {b c d : G₀} : (a / b) / (c / d) = (a * d) / (b * c) := -by rw [div_div_eq_mul_div, div_mul_eq_mul_div, div_div_eq_div_mul] +division_comm_monoid.div_div_div_eq _ _ _ _ lemma div_mul_eq_div_mul_one_div (a b c : G₀) : a / (b * c) = (a / b) * (1 / c) := -by rw [← div_div_eq_div_mul, ← div_eq_mul_one_div] +division_comm_monoid.div_mul_eq_div_mul_one_div _ _ _ lemma div_helper {a : G₀} (b : G₀) (h : a ≠ 0) : (1 / (a * b)) * a = 1 / b := by rw [div_mul_eq_mul_div, one_mul, div_mul_right _ h] @@ -1043,23 +1012,22 @@ end comm_group_with_zero section comm_group_with_zero variables [comm_group_with_zero G₀] {a b c d : G₀} -lemma div_eq_inv_mul : a / b = b⁻¹ * a := -by rw [div_eq_mul_inv, mul_comm] +lemma div_eq_inv_mul : a / b = b⁻¹ * a := division_comm_monoid.div_eq_inv_mul _ _ lemma mul_div_right_comm (a b c : G₀) : (a * b) / c = (a / c) * b := -by rw [div_eq_mul_inv, mul_assoc, mul_comm b, ← mul_assoc, div_eq_mul_inv] +division_comm_monoid.mul_div_right_comm _ _ _ -lemma mul_comm_div' (a b c : G₀) : (a / b) * c = a * (c / b) := -by rw [← mul_div_assoc, mul_div_right_comm] +lemma mul_comm_div' (a b c : G₀) : a / b * c = a * (c / b) := +division_comm_monoid.mul_comm_div _ _ _ lemma div_mul_comm' (a b c : G₀) : (a / b) * c = (c / b) * a := -by rw [div_mul_eq_mul_div, mul_comm, mul_div_right_comm] +division_comm_monoid.div_mul_comm _ _ _ lemma mul_div_comm (a b c : G₀) : a * (b / c) = b * (a / c) := -by rw [← mul_div_assoc, mul_comm, mul_div_assoc] +division_comm_monoid.mul_div_left_comm _ _ _ lemma div_right_comm (a : G₀) : (a / b) / c = (a / c) / b := -by rw [div_div_eq_div_mul, div_div_eq_div_mul, mul_comm] +division_comm_monoid.div_right_comm _ _ _ @[field_simps] lemma div_eq_div_iff (hb : b ≠ 0) (hd : d ≠ 0) : a / b = c / d ↔ a * d = c * b := calc a / b = c / d ↔ a / b * (b * d) = c / d * (b * d) : diff --git a/src/analysis/special_functions/trigonometric/complex.lean b/src/analysis/special_functions/trigonometric/complex.lean index d73e52587085e..0f877132d3d2b 100644 --- a/src/analysis/special_functions/trigonometric/complex.lean +++ b/src/analysis/special_functions/trigonometric/complex.lean @@ -217,7 +217,7 @@ by simpa [mul_comm x] using strict_concave_on_sin_Icc.concave_on.2 ⟨le_rfl, pi lemma mul_lt_sin {x : ℝ} (hx : 0 < x) (hx' : x < π / 2) : (2 / π) * x < sin x := begin rw [←inv_div], - simpa [pi_div_two_pos.ne', mul_nonneg, inv_nonneg] using @lt_sin_mul ((π / 2)⁻¹ * x) _ _, + simpa [-division_monoid.inv_div, pi_div_two_pos.ne'] using @lt_sin_mul ((π / 2)⁻¹ * x) _ _, { exact mul_pos (inv_pos.2 pi_div_two_pos) hx }, { rwa [←div_eq_inv_mul, div_lt_one pi_div_two_pos] }, end @@ -227,7 +227,7 @@ of Jordan's inequality, the other half is `real.sin_lt` -/ lemma mul_le_sin {x : ℝ} (hx : 0 ≤ x) (hx' : x ≤ π / 2) : (2 / π) * x ≤ sin x := begin rw [←inv_div], - simpa [pi_div_two_pos.ne', mul_nonneg, inv_nonneg] using @le_sin_mul ((π / 2)⁻¹ * x) _ _, + simpa [-division_monoid.inv_div, pi_div_two_pos.ne'] using @le_sin_mul ((π / 2)⁻¹ * x) _ _, { exact mul_nonneg (inv_nonneg.2 pi_div_two_pos.le) hx }, { rwa [←div_eq_inv_mul, div_le_one pi_div_two_pos] }, end From 1e8f3817346bffaae9675f9a0d4e7a1cbd21ff9e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Mon, 9 May 2022 02:37:42 +0000 Subject: [PATCH 28/37] feat(algebraic_topology/dold_kan): defining some null homotopic maps (#13085) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- docs/references.bib | 30 +++ src/algebra/homology/complex_shape.lean | 7 + .../dold_kan/homotopies.lean | 205 ++++++++++++++++++ .../dold_kan/notations.lean | 20 ++ 4 files changed, 262 insertions(+) create mode 100644 src/algebraic_topology/dold_kan/homotopies.lean create mode 100644 src/algebraic_topology/dold_kan/notations.lean diff --git a/docs/references.bib b/docs/references.bib index 180e8c9d71c18..35eae3137cecd 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -421,6 +421,22 @@ @InProceedings{ deligne_formulaire mrreviewer = {Jacques Velu} } +@Article{ dold1958, + author = {Dold, Albrecht}, + title = {Homology of symmetric products and other functors of + complexes}, + journal = {Ann. of Math. (2)}, + fjournal = {Annals of Mathematics. Second Series}, + volume = {68}, + year = {1958}, + pages = {54--80}, + issn = {0003-486X}, + mrclass = {55.00}, + mrnumber = {97057}, + mrreviewer = {Sze-tsen Hu}, + doi = {10.2307/1970043} +} + @Misc{ dupuis-lewis-macbeth2022, title = {Formalized functional analysis with semilinear maps}, author = {Frédéric Dupuis and Robert Y. Lewis and Heather @@ -657,6 +673,20 @@ @Book{ GierzEtAl1980 mrreviewer = {James W. Lea, Jr.} } +@Book{ goerss-jardine-2009, + author = {Goerss, Paul G. and Jardine, John F.}, + title = {Simplicial homotopy theory}, + series = {Modern Birkh\"{a}user Classics}, + note = {Reprint of the 1999 edition [MR1711612]}, + publisher = {Birkh\"{a}user Verlag, Basel}, + year = {2009}, + pages = {xvi+510}, + isbn = {978-3-0346-0188-7}, + mrclass = {55U10 (18G55)}, + mrnumber = {2840650}, + doi = {10.1007/978-3-0346-0189-4} +} + @Book{ Gordon55, author = {Russel A. Gordon}, title = {The integrals of Lebesgue, Denjoy, Perron, and Henstock}, diff --git a/src/algebra/homology/complex_shape.lean b/src/algebra/homology/complex_shape.lean index 8be04efb55ad3..1a31f498d7581 100644 --- a/src/algebra/homology/complex_shape.lean +++ b/src/algebra/homology/complex_shape.lean @@ -168,6 +168,9 @@ def down' {α : Type*} [add_right_cancel_semigroup α] (a : α) : complex_shape next_eq := λ i j k hi hj, add_right_cancel (hi.trans (hj.symm)), prev_eq := λ i j k hi hj, hi.symm.trans hj, } +lemma down'_mk {α : Type*} [add_right_cancel_semigroup α] (a : α) + (i j : α) (h : j + a = i) : (down' a).rel i j := h + /-- The `complex_shape` appropriate for cohomology, so `d : X i ⟶ X j` only when `j = i + 1`. -/ @@ -182,4 +185,8 @@ The `complex_shape` appropriate for homology, so `d : X i ⟶ X j` only when `i def down (α : Type*) [add_right_cancel_semigroup α] [has_one α] : complex_shape α := down' 1 +lemma down_mk {α : Type*} [add_right_cancel_semigroup α] [has_one α] + (i j : α) (h : j + 1 = i) : (down α).rel i j := +down'_mk (1 : α) i j h + end complex_shape diff --git a/src/algebraic_topology/dold_kan/homotopies.lean b/src/algebraic_topology/dold_kan/homotopies.lean new file mode 100644 index 0000000000000..631f3cb9fe164 --- /dev/null +++ b/src/algebraic_topology/dold_kan/homotopies.lean @@ -0,0 +1,205 @@ +/- +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.homotopy +import algebraic_topology.dold_kan.notations + +/-! + +# Construction of homotopies for the Dold-Kan correspondence + +TODO (@joelriou) continue adding the various files references below + +(The general strategy of proof of the Dold-Kan correspondence is explained +in `equivalence.lean`.) + +The purpose of the files `homotopies.lean`, `faces.lean`, `projections.lean` +and `p_infty.lean` is to construct an idempotent endomorphism +`P_infty : K[X] ⟶ K[X]` of the alternating face map complex +for each `X : simplicial_object C` when `C` is a preadditive category. +In the case `C` is abelian, this `P_infty` shall be the projection on the +normalized Moore subcomplex of `K[X]` associated to the decomposition of the +complex `K[X]` as a direct sum of this normalized subcomplex and of the +degenerate subcomplex. + +In `p_infty.lean`, this endomorphism `P_infty` shall be obtained by +passing to the limit idempotent endomorphisms `P q` for all `(q : ℕ)`. +These endomorphisms `P q` are defined by induction. The idea is to +start from the identity endomorphism `P 0` of `K[X]` and to ensure by +induction that the `q` higher face maps (except $d_0$) vanish on the +image of `P q`. Then, in a certain degree `n`, the image of `P q` for +a big enough `q` will be contained in the normalized subcomplex. This +construction is done in `projections.lean`. + +It would be easy to define the `P q` degreewise (similarly as it is done +in *Simplicial Homotopy Theory* by Goerrs-Jardine p. 149), but then we would +have to prove that they are compatible with the differential (i.e. they +are chain complex maps), and also that they are homotopic to the identity. +These two verifications are quite technical. In order to reduce the number +of such technical lemmas, the strategy that is followed here is to define +a series of null homotopic maps `Hσ q` (attached to families of maps `hσ`) +and use these in order to construct `P q` : the endomorphisms `P q` +shall basically be obtained by altering the identity endomorphism by adding +null homotopic maps, so that we get for free that they are morphisms +of chain complexes and that they are homotopic to the identity. The most +technical verifications that are needed about the null homotopic maps `Hσ` +are obtained in `faces.lean`. + +In this file `homotopies.lean`, we define the null homotopic maps +`Hσ q : K[X] ⟶ K[X]`, show that they are natural (see `nat_trans_Hσ`) and +compatible the application of additive functors (see `map_Hσ`). + +## References +* [Albrecht Dold, *Homology of Symmetric Products and Other Functors of Complexes*][dold1958] +* [Paul G. Goerss, John F. Jardine, *Simplical Homotopy Theory*][goerss-jardine-2009] + +-/ + +open category_theory +open category_theory.category +open category_theory.limits +open category_theory.preadditive +open category_theory.simplicial_object +open homotopy +open opposite +open_locale simplicial dold_kan + +noncomputable theory + +namespace algebraic_topology + +namespace dold_kan + +variables {C : Type*} [category C] [preadditive C] +variables {X : simplicial_object C} + +/-- As we are using chain complexes indexed by `ℕ`, we shall need the relation +`c` such `c m n` if and only if `n+1=m`. -/ +abbreviation c := complex_shape.down ℕ + +/-- Helper when we need some `c.rel i j` (i.e. `complex_shape.down ℕ`), +e.g. `c_mk n (n+1) rfl` -/ +lemma c_mk (i j : ℕ) (h : j+1 = i) : c.rel i j := +complex_shape.down_mk i j h + +/-- This lemma is meant to be used with `null_homotopic_map'_f_of_not_rel_left` -/ +lemma cs_down_0_not_rel_left (j : ℕ) : ¬c.rel 0 j := +begin + intro hj, + dsimp at hj, + apply nat.not_succ_le_zero j, + rw [nat.succ_eq_add_one, hj], +end + +/-- The sequence of maps which gives the null homotopic maps `Hσ` that shall be in +the inductive construction of the projections `P q : K[X] ⟶ K[X]` -/ +def hσ (q : ℕ) (n : ℕ) : X _[n] ⟶ X _[n+1] := +if n Date: Mon, 9 May 2022 02:37:43 +0000 Subject: [PATCH 29/37] feat(set_theory/pgame/birthday): Birthdays of ordinals (#13714) --- src/set_theory/game/birthday.lean | 45 ++++++++++++++++++++++++------- src/set_theory/game/ordinal.lean | 1 - 2 files changed, 36 insertions(+), 10 deletions(-) diff --git a/src/set_theory/game/birthday.lean b/src/set_theory/game/birthday.lean index 6cef0d34fa933..f4a5d013d524e 100644 --- a/src/set_theory/game/birthday.lean +++ b/src/set_theory/game/birthday.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Violeta Hernández Palacios -/ +import set_theory.game.ordinal import set_theory.ordinal.arithmetic -import set_theory.game.pgame /-! # Birthdays of games @@ -26,6 +26,8 @@ prove the basic properties about these. universe u +open ordinal + namespace pgame /-- The birthday of a pre-game is inductively defined as the least strict upper bound of the @@ -33,20 +35,20 @@ birthdays of its left and right games. It may be thought as the "step" in which constructed. -/ noncomputable def birthday : pgame.{u} → ordinal.{u} | ⟨xl, xr, xL, xR⟩ := - max (ordinal.lsub.{u u} $ λ i, birthday (xL i)) (ordinal.lsub.{u u} $ λ i, birthday (xR i)) + max (lsub.{u u} $ λ i, birthday (xL i)) (lsub.{u u} $ λ i, birthday (xR i)) theorem birthday_def (x : pgame) : birthday x = max - (ordinal.lsub.{u u} (λ i, birthday (x.move_left i))) - (ordinal.lsub.{u u} (λ i, birthday (x.move_right i))) := + (lsub.{u u} (λ i, birthday (x.move_left i))) + (lsub.{u u} (λ i, birthday (x.move_right i))) := by { cases x, rw birthday, refl } theorem birthday_move_left_lt {x : pgame} (i : x.left_moves) : (x.move_left i).birthday < x.birthday := -by { cases x, rw birthday, exact lt_max_of_lt_left (ordinal.lt_lsub _ i) } +by { cases x, rw birthday, exact lt_max_of_lt_left (lt_lsub _ i) } theorem birthday_move_right_lt {x : pgame} (i : x.right_moves) : (x.move_right i).birthday < x.birthday := -by { cases x, rw birthday, exact lt_max_of_lt_right (ordinal.lt_lsub _ i) } +by { cases x, rw birthday, exact lt_max_of_lt_right (lt_lsub _ i) } theorem lt_birthday_iff {x : pgame} {o : ordinal} : o < x.birthday ↔ (∃ i : x.left_moves, o ≤ (x.move_left i).birthday) ∨ @@ -57,9 +59,9 @@ begin intro h, cases lt_max_iff.1 h with h' h', { left, - rwa ordinal.lt_lsub_iff at h' }, + rwa lt_lsub_iff at h' }, { right, - rwa ordinal.lt_lsub_iff at h' } }, + rwa lt_lsub_iff at h' } }, { rintro (⟨i, hi⟩ | ⟨i, hi⟩), { exact hi.trans_lt (birthday_move_left_lt i) }, { exact hi.trans_lt (birthday_move_right_lt i) } } @@ -70,7 +72,7 @@ theorem relabelling.birthday_congr : ∀ {x y : pgame.{u}}, relabelling x y → rw [birthday, birthday], congr' 1, all_goals - { apply ordinal.lsub_eq_of_range_eq.{u u u}, + { apply lsub_eq_of_range_eq.{u u u}, ext i, split }, { rintro ⟨j, rfl⟩, @@ -109,4 +111,29 @@ begin exact max_bot_right 1 end +@[simp] theorem neg_birthday : ∀ x : pgame, (-x).birthday = x.birthday +| ⟨xl, xr, xL, xR⟩ := begin + rw [birthday_def, birthday_def, max_comm], + congr; funext; apply neg_birthday +end + +@[simp] theorem to_pgame_birthday (o : ordinal) : o.to_pgame.birthday = o := +begin + induction o using ordinal.induction with o IH, + rw [to_pgame_def, pgame.birthday], + convert max_eq_left_iff.2 (ordinal.zero_le _), + { apply lsub_empty }, + { nth_rewrite 0 ←lsub_typein o, + congr, + exact funext (λ x, (IH _ (typein_lt_self x)).symm) } +end + +theorem le_birthday : ∀ x : pgame, x ≤ x.birthday.to_pgame +| ⟨xl, _, xL, _⟩ := +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, le_iff_neg_ge, neg_neg] at h + end pgame diff --git a/src/set_theory/game/ordinal.lean b/src/set_theory/game/ordinal.lean index 8d8750d01ee48..93409d6988639 100644 --- a/src/set_theory/game/ordinal.lean +++ b/src/set_theory/game/ordinal.lean @@ -21,7 +21,6 @@ set consists of all previous ordinals. # Todo - Extend this map to `game` and `surreal`. -- Prove that `birthday o.to_pgame = o`. -/ local infix ` ≈ ` := pgame.equiv From 5253153299d6da8255c23c69fa19c24cef88e8d1 Mon Sep 17 00:00:00 2001 From: antoinelab01 Date: Mon, 9 May 2022 02:37:44 +0000 Subject: [PATCH 30/37] feat(algebra/category/Module/basic): `iso.hom_congr`agrees with `linear_equiv.arrow_congr` (#13954) Co-authored-by: antoinelab01 <66086247+antoinelab01@users.noreply.github.com> --- src/algebra/category/Module/basic.lean | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/src/algebra/category/Module/basic.lean b/src/algebra/category/Module/basic.lean index 4d93857824f82..f5bc120c5a0aa 100644 --- a/src/algebra/category/Module/basic.lean +++ b/src/algebra/category/Module/basic.lean @@ -7,6 +7,7 @@ import algebra.category.Group.basic import category_theory.limits.shapes.kernels import category_theory.linear import linear_algebra.basic +import category_theory.conj /-! # The category of `R`-modules @@ -263,6 +264,14 @@ instance : linear S (Module.{v} S) := smul_comp' := by { intros, ext, simp }, comp_smul' := by { intros, ext, simp }, } +variables {X Y X' Y' : Module.{v} S} + +lemma iso.hom_congr_eq_arrow_congr (i : X ≅ X') (j : Y ≅ Y') (f : X ⟶ Y) : + iso.hom_congr i j f = linear_equiv.arrow_congr i.to_linear_equiv j.to_linear_equiv f := rfl + +lemma iso.conj_eq_conj (i : X ≅ X') (f : End X) : + iso.conj i f = linear_equiv.conj i.to_linear_equiv f := rfl + end end Module From 34b61e3d612e7c44a42902abf73473e662bc4c84 Mon Sep 17 00:00:00 2001 From: Eric Rodriguez Date: Mon, 9 May 2022 02:37:45 +0000 Subject: [PATCH 31/37] chore(algebra/regular/*): generalisation linter (#13955) --- src/algebra/regular/basic.lean | 18 +++++--- src/algebra/regular/smul.lean | 79 ++++++++++++++++++++-------------- 2 files changed, 58 insertions(+), 39 deletions(-) diff --git a/src/algebra/regular/basic.lean b/src/algebra/regular/basic.lean index 75f5a84ab02a1..86bd53f1901de 100644 --- a/src/algebra/regular/basic.lean +++ b/src/algebra/regular/basic.lean @@ -241,6 +241,18 @@ lemma not_is_regular_zero [nontrivial R] : ¬ is_regular (0 : R) := end mul_zero_class +section mul_one_class + +variable [mul_one_class R] + +/-- If multiplying by `1` on either side is the identity, `1` is regular. -/ +@[to_additive "If adding `0` on either side is the identity, `0` is regular."] +lemma is_regular_one : is_regular (1 : R) := +⟨λ a b ab, (one_mul a).symm.trans (eq.trans ab (one_mul b)), + λ a b ab, (mul_one a).symm.trans (eq.trans ab (mul_one b))⟩ + +end mul_one_class + section comm_semigroup variable [comm_semigroup R] @@ -259,12 +271,6 @@ section monoid variables [monoid R] -/-- In a monoid, `1` is regular. -/ -@[to_additive "In an additive monoid, `0` is regular."] -lemma is_regular_one : is_regular (1 : R) := -⟨λ a b ab, (one_mul a).symm.trans (eq.trans ab (one_mul b)), - λ a b ab, (mul_one a).symm.trans (eq.trans ab (mul_one b))⟩ - /-- 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 := diff --git a/src/algebra/regular/smul.lean b/src/algebra/regular/smul.lean index b79b53161f311..b45087acf3d3a 100644 --- a/src/algebra/regular/smul.lean +++ b/src/algebra/regular/smul.lean @@ -5,6 +5,7 @@ Authors: Damiano Testa -/ import algebra.smul_with_zero import algebra.regular.basic + /-! # Action of regular elements on a module @@ -76,35 +77,22 @@ lemma is_left_regular [has_mul R] {a : R} (h : is_smul_regular R a) : lemma is_right_regular [has_mul R] {a : R} (h : is_smul_regular R (mul_opposite.op a)) : is_right_regular a := h -end has_scalar - -section monoid - -variables [monoid R] [mul_action R M] - -variable (M) - -/-- One is `M`-regular always. -/ -@[simp] lemma one : is_smul_regular M (1 : R) := -λ a b ab, by rwa [one_smul, one_smul] at ab - -variable {M} - -lemma mul (ra : is_smul_regular M a) (rb : is_smul_regular M b) : - is_smul_regular M (a * b) := +lemma mul [has_mul R] [is_scalar_tower R R M] + (ra : is_smul_regular M a) (rb : is_smul_regular M b) : is_smul_regular M (a * b) := ra.smul rb -lemma of_mul (ab : is_smul_regular M (a * b)) : +lemma of_mul [has_mul R] [is_scalar_tower R R M] (ab : is_smul_regular M (a * b)) : is_smul_regular M b := by { rw ← smul_eq_mul at ab, exact ab.of_smul _ } -@[simp] lemma mul_iff_right (ha : is_smul_regular M a) : +@[simp] lemma mul_iff_right [has_mul R] [is_scalar_tower R R M] (ha : is_smul_regular M a) : is_smul_regular M (a * b) ↔ is_smul_regular M b := ⟨of_mul, ha.mul⟩ /-- Two elements `a` and `b` are `M`-regular if and only if both products `a * b` and `b * a` are `M`-regular. -/ -lemma mul_and_mul_iff : is_smul_regular M (a * b) ∧ is_smul_regular M (b * a) ↔ +lemma mul_and_mul_iff [has_mul R] [is_scalar_tower R R M] : + is_smul_regular M (a * b) ∧ is_smul_regular M (b * a) ↔ is_smul_regular M a ∧ is_smul_regular M b := begin refine ⟨_, _⟩, @@ -114,6 +102,24 @@ begin exact ⟨ha.mul hb, hb.mul ha⟩ } end +end has_scalar + +section monoid + +variables [monoid R] [mul_action R M] + +variable (M) + +/-- One is `M`-regular always. -/ +@[simp] lemma one : is_smul_regular M (1 : R) := +λ a b ab, by rwa [one_smul, one_smul] at ab + +variable {M} + +/-- An element of `R` admitting a left inverse is `M`-regular. -/ +lemma of_mul_eq_one (h : a * b = 1) : is_smul_regular M b := +of_mul (by { rw h, exact one M }) + /-- Any power of an `M`-regular element is `M`-regular. -/ lemma pow (n : ℕ) (ra : is_smul_regular M a) : is_smul_regular M (a ^ n) := begin @@ -133,10 +139,21 @@ end end monoid +section monoid_smul + +variables [monoid S] [has_scalar R M] [has_scalar R S] [mul_action S M] [is_scalar_tower R S M] + +/-- An element of `S` admitting a left inverse in `R` is `M`-regular. -/ +lemma of_smul_eq_one (h : a • s = 1) : is_smul_regular M s := +of_smul a (by { rw h, exact one M }) + +end monoid_smul + section monoid_with_zero -variables [monoid_with_zero R] [monoid_with_zero S] [has_zero M] [mul_action_with_zero R M] - [mul_action_with_zero R S] [mul_action_with_zero S M] [is_scalar_tower R S M] +variables [monoid_with_zero R] [monoid_with_zero S] [has_zero M] + [mul_action_with_zero R M] [mul_action_with_zero R S] [mul_action_with_zero S M] + [is_scalar_tower R S M] /-- The element `0` is `M`-regular if and only if `M` is trivial. -/ protected lemma subsingleton (h : is_smul_regular M (0 : R)) : subsingleton M := @@ -162,19 +179,11 @@ zero_iff_subsingleton.mpr sM lemma not_zero [nM : nontrivial M] : ¬ is_smul_regular M (0 : R) := not_zero_iff.mpr nM -/-- An element of `S` admitting a left inverse in `R` is `M`-regular. -/ -lemma of_smul_eq_one (h : a • s = 1) : is_smul_regular M s := -of_smul a (by { rw h, exact one M }) - -/-- An element of `R` admitting a left inverse is `M`-regular. -/ -lemma of_mul_eq_one (h : a * b = 1) : is_smul_regular M b := -of_mul (by { rw h, exact one M }) - end monoid_with_zero -section comm_monoid +section comm_semigroup -variables [comm_monoid R] [mul_action R M] +variables [comm_semigroup R] [has_scalar R M] [is_scalar_tower R R M] /-- A product is `M`-regular if and only if the factors are. -/ lemma mul_iff : is_smul_regular M (a * b) ↔ @@ -184,7 +193,7 @@ begin exact ⟨λ ab, ⟨ab, by rwa mul_comm⟩, λ rab, rab.1⟩ end -end comm_monoid +end comm_semigroup end is_smul_regular @@ -203,7 +212,9 @@ end end group -variables [monoid_with_zero R] [has_zero M] [mul_action_with_zero R M] +section units + +variables [monoid R] [mul_action R M] /-- Any element in `Rˣ` is `M`-regular. -/ lemma units.is_smul_regular (a : Rˣ) : is_smul_regular M (a : R) := @@ -215,3 +226,5 @@ begin rcases ua with ⟨a, rfl⟩, exact a.is_smul_regular M end + +end units From afec1d73edeea4f075119daef50c08aedde57ff7 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 9 May 2022 02:37:47 +0000 Subject: [PATCH 32/37] fix(tactics/alias): Make docstring calculation available to to_additive (#13968) PR #13944 fixed the docstrings for iff-style aliases, but because of code duplication I added in #13330 this did not apply to aliases introduced by `to_additive`. This fixes that. --- src/algebra/group/to_additive.lean | 5 ++-- src/tactic/alias.lean | 44 ++++++++++++++++++++++-------- 2 files changed, 36 insertions(+), 13 deletions(-) diff --git a/src/algebra/group/to_additive.lean b/src/algebra/group/to_additive.lean index 1b16d6cb7ad9f..cf17f07b07425 100644 --- a/src/algebra/group/to_additive.lean +++ b/src/algebra/group/to_additive.lean @@ -553,9 +553,10 @@ protected meta def attr : user_attribute unit value_type := match val.doc with | some doc := add_doc_string tgt doc | none := do - some alias_name ← tactic.alias.get_alias_target src | skip, + some alias_target ← tactic.alias.get_alias_target src | skip, + let alias_name := alias_target.to_name, some add_alias_name ← pure (dict.find alias_name) | skip, - add_doc_string tgt ("**Alias** of `" ++ to_string add_alias_name ++ "`.") + add_doc_string tgt alias_target.to_string end } add_tactic_doc diff --git a/src/tactic/alias.lean b/src/tactic/alias.lean index 0ade90b333a43..a4fe87f82de47 100644 --- a/src/tactic/alias.lean +++ b/src/tactic/alias.lean @@ -46,10 +46,29 @@ open lean.parser tactic interactive namespace tactic.alias -@[user_attribute] meta def alias_attr : user_attribute unit name := +/-- An alias can be in one of three forms -/ +@[derive has_reflect] +meta inductive target +| plain : name -> target +| forward : name -> target +| backwards : name -> target + +/-- The name underlying an alias target -/ +meta def target.to_name : target → name +| (target.plain n) := n +| (target.forward n) := n +| (target.backwards n) := n + +/-- The docstring for an alias. Used by `alias` _and_ by `to_additive` -/ +meta def target.to_string : target → string +| (target.plain n) := sformat!"**Alias** of {n}`." +| (target.forward n) := sformat!"**Alias** of the forward direction of {n}`." +| (target.backwards n) := sformat!"**Alias** of the reverse direction of {n}`." + +@[user_attribute] meta def alias_attr : user_attribute unit target := { name := `alias, descr := "This definition is an alias of another.", parser := failed } -meta def alias_direct (d : declaration) (doc : string) (al : name) : tactic unit := +meta def alias_direct (d : declaration) (al : name) : tactic unit := do updateex_env $ λ env, env.add (match d.to_definition with | declaration.defn n ls t _ _ _ := @@ -59,23 +78,26 @@ do updateex_env $ λ env, declaration.thm al ls t $ task.pure $ expr.const n (level.param <$> ls) | _ := undefined end), - alias_attr.set al d.to_name tt, - add_doc_string al doc + let target := target.plain d.to_name, + alias_attr.set al target tt, + add_doc_string al target.to_string meta def mk_iff_mp_app (iffmp : name) : expr → (ℕ → expr) → tactic expr | (expr.pi n bi e t) f := expr.lam n bi e <$> mk_iff_mp_app t (λ n, f (n+1) (expr.var n)) | `(%%a ↔ %%b) f := pure $ @expr.const tt iffmp [] a b (f 0) | _ f := fail "Target theorem must have the form `Π x y z, a ↔ b`" -meta def alias_iff (d : declaration) (doc : string) (al : name) (iffmp : name) : tactic unit := +meta def alias_iff (d : declaration) (al : name) (is_forward : bool) : tactic unit := (if al = `_ then skip else get_decl al >> skip) <|> do let ls := d.univ_params, let t := d.type, + let target := if is_forward then target.forward d.to_name else target.backwards d.to_name, + let iffmp := if is_forward then `iff.mp else `iff.mpr, v ← mk_iff_mp_app iffmp t (λ_, expr.const d.to_name (level.param <$> ls)), t' ← infer_type v, updateex_env $ λ env, env.add (declaration.thm al ls t' $ task.pure v), - alias_attr.set al d.to_name tt, - add_doc_string al doc + alias_attr.set al target tt, + add_doc_string al target.to_string meta def make_left_right : name → tactic (name × name) | (name.mk_string s p) := do @@ -134,15 +156,15 @@ do old ← ident, do { tk "←" <|> tk "<-", aliases ← many ident, - ↑(aliases.mmap' $ λ al, alias_direct d (doc al "") al) } <|> + ↑(aliases.mmap' $ λ al, alias_direct d al) } <|> do { tk "↔" <|> tk "<->", (left, right) ← mcond ((tk ".." >> pure tt) <|> pure ff) (make_left_right old <|> fail "invalid name for automatic name generation") (prod.mk <$> types.ident_ <*> types.ident_), - alias_iff d (doc left "the forward direction of ") left `iff.mp, - alias_iff d (doc right "the reverse direction of ") right `iff.mpr } + alias_iff d left tt, + alias_iff d right ff } add_tactic_doc { name := "alias", @@ -154,7 +176,7 @@ meta def get_lambda_body : expr → expr | (expr.lam _ _ _ b) := get_lambda_body b | a := a -meta def get_alias_target (n : name) : tactic (option name) := +meta def get_alias_target (n : name) : tactic (option target) := do tt ← has_attribute' `alias n | pure none, v ← alias_attr.get_param n, pure $ some v From 594ceda90604bfc6e77a852a9d68ae78ea13cdfc Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 9 May 2022 02:37:47 +0000 Subject: [PATCH 33/37] feat(analysis/normed_space/exponential): Generalize `field` lemmas to `division_ring` (#13997) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This generalizes the lemmas about `exp 𝕂 𝕂` to lemmas about `exp 𝕂 𝔸` where `𝔸` is a `division_ring`. This moves the lemmas down to the appropriate division_ring sections, and replaces the word `field` with `div` in their names, since `division_ring` is a mouthful, and really the name reflects the use of `/` in the definition. --- src/analysis/normed_space/exponential.lean | 115 ++++++++++-------- .../special_functions/exponential.lean | 6 +- src/analysis/specific_limits/normed.lean | 2 +- .../derangements/exponential.lean | 2 +- 4 files changed, 67 insertions(+), 58 deletions(-) diff --git a/src/analysis/normed_space/exponential.lean b/src/analysis/normed_space/exponential.lean index 870ab76e23e08..6fab164c2a7b2 100644 --- a/src/analysis/normed_space/exponential.lean +++ b/src/analysis/normed_space/exponential.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2021 Anatole Dedecker. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Anatole Dedecker +Authors: Anatole Dedecker, Eric Wieser -/ import analysis.specific_limits.basic import analysis.analytic.basic @@ -88,31 +88,12 @@ lemma exp_series_apply_eq' (x : 𝔸) : (λ n, exp_series 𝕂 𝔸 n (λ _, x)) = (λ n, (n!⁻¹ : 𝕂) • x^n) := funext (exp_series_apply_eq x) -lemma exp_series_apply_eq_field [topological_space 𝕂] [topological_ring 𝕂] (x : 𝕂) (n : ℕ) : - exp_series 𝕂 𝕂 n (λ _, x) = x^n / n! := -begin - rw [div_eq_inv_mul, ←smul_eq_mul], - exact exp_series_apply_eq x n, -end - -lemma exp_series_apply_eq_field' [topological_space 𝕂] [topological_ring 𝕂] (x : 𝕂) : - (λ n, exp_series 𝕂 𝕂 n (λ _, x)) = (λ n, x^n / n!) := -funext (exp_series_apply_eq_field x) - lemma exp_series_sum_eq (x : 𝔸) : (exp_series 𝕂 𝔸).sum x = ∑' (n : ℕ), (n!⁻¹ : 𝕂) • x^n := tsum_congr (λ n, exp_series_apply_eq x n) -lemma exp_series_sum_eq_field [topological_space 𝕂] [topological_ring 𝕂] (x : 𝕂) : - (exp_series 𝕂 𝕂).sum x = ∑' (n : ℕ), x^n / n! := -tsum_congr (λ n, exp_series_apply_eq_field x n) - lemma exp_eq_tsum : exp 𝕂 = (λ x : 𝔸, ∑' (n : ℕ), (n!⁻¹ : 𝕂) • x^n) := funext exp_series_sum_eq -lemma exp_eq_tsum_field [topological_space 𝕂] [topological_ring 𝕂] : - exp 𝕂 = (λ x : 𝕂, ∑' (n : ℕ), x^n / n!) := -funext exp_series_sum_eq_field - @[simp] lemma exp_zero [t2_space 𝔸] : exp 𝕂 (0 : 𝔸) = 1 := begin suffices : (λ x : 𝔸, ∑' (n : ℕ), (n!⁻¹ : 𝕂) • x^n) 0 = ∑' (n : ℕ), if n = 0 then 1 else 0, @@ -145,6 +126,25 @@ lemma commute.exp [t2_space 𝔸] {x y : 𝔸} (h : commute x y) : commute (exp end topological_algebra +section topological_division_algebra +variables {𝕂 𝔸 : Type*} [field 𝕂] [division_ring 𝔸] [algebra 𝕂 𝔸] [topological_space 𝔸] + [topological_ring 𝔸] + +lemma exp_series_apply_eq_div (x : 𝔸) (n : ℕ) : exp_series 𝕂 𝔸 n (λ _, x) = x^n / n! := +by rw [div_eq_mul_inv, ←(nat.cast_commute n! (x ^ n)).inv_left₀.eq, ←smul_eq_mul, + exp_series_apply_eq, inv_nat_cast_smul_eq _ _ _ _] + +lemma exp_series_apply_eq_div' (x : 𝔸) : (λ n, exp_series 𝕂 𝔸 n (λ _, x)) = (λ n, x^n / n!) := +funext (exp_series_apply_eq_div x) + +lemma exp_series_sum_eq_div (x : 𝔸) : (exp_series 𝕂 𝔸).sum x = ∑' (n : ℕ), x^n / n! := +tsum_congr (exp_series_apply_eq_div x) + +lemma exp_eq_tsum_div : exp 𝕂 = (λ x : 𝔸, ∑' (n : ℕ), x^n / n!) := +funext exp_series_sum_eq_div + +end topological_division_algebra + section normed section any_field_any_algebra @@ -166,15 +166,6 @@ begin exact norm_exp_series_summable_of_mem_ball x hx end -lemma norm_exp_series_field_summable_of_mem_ball (x : 𝕂) - (hx : x ∈ emetric.ball (0 : 𝕂) (exp_series 𝕂 𝕂).radius) : - summable (λ n, ∥x^n / n!∥) := -begin - change summable (norm ∘ _), - rw ← exp_series_apply_eq_field', - exact norm_exp_series_summable_of_mem_ball x hx -end - section complete_algebra variables [complete_space 𝔸] @@ -189,10 +180,6 @@ lemma exp_series_summable_of_mem_ball' (x : 𝔸) summable (λ n, (n!⁻¹ : 𝕂) • x^n) := summable_of_summable_norm (norm_exp_series_summable_of_mem_ball' x hx) -lemma exp_series_field_summable_of_mem_ball [complete_space 𝕂] (x : 𝕂) - (hx : x ∈ emetric.ball (0 : 𝕂) (exp_series 𝕂 𝕂).radius) : summable (λ n, x^n / n!) := -summable_of_summable_norm (norm_exp_series_field_summable_of_mem_ball x hx) - lemma exp_series_has_sum_exp_of_mem_ball (x : 𝔸) (hx : x ∈ emetric.ball (0 : 𝔸) (exp_series 𝕂 𝔸).radius) : has_sum (λ n, exp_series 𝕂 𝔸 n (λ _, x)) (exp 𝕂 x) := @@ -206,13 +193,6 @@ begin exact exp_series_has_sum_exp_of_mem_ball x hx end -lemma exp_series_field_has_sum_exp_of_mem_ball [complete_space 𝕂] (x : 𝕂) - (hx : x ∈ emetric.ball (0 : 𝕂) (exp_series 𝕂 𝕂).radius) : has_sum (λ n, x^n / n!) (exp 𝕂 x) := -begin - rw ← exp_series_apply_eq_field', - exact exp_series_has_sum_exp_of_mem_ball x hx -end - lemma has_fpower_series_on_ball_exp_of_radius_pos (h : 0 < (exp_series 𝕂 𝔸).radius) : has_fpower_series_on_ball (exp 𝕂) (exp_series 𝕂 𝔸) 0 (exp_series 𝕂 𝔸).radius := (exp_series 𝕂 𝔸).has_fpower_series_on_ball h @@ -305,6 +285,30 @@ section any_field_division_algebra variables {𝕂 𝔸 : Type*} [nondiscrete_normed_field 𝕂] [normed_division_ring 𝔸] [normed_algebra 𝕂 𝔸] +variables (𝕂) + +lemma norm_exp_series_div_summable_of_mem_ball (x : 𝔸) + (hx : x ∈ emetric.ball (0 : 𝔸) (exp_series 𝕂 𝔸).radius) : + summable (λ n, ∥x^n / n!∥) := +begin + change summable (norm ∘ _), + rw ← exp_series_apply_eq_div' x, + exact norm_exp_series_summable_of_mem_ball x hx +end + +lemma exp_series_div_summable_of_mem_ball [complete_space 𝔸] (x : 𝔸) + (hx : x ∈ emetric.ball (0 : 𝔸) (exp_series 𝕂 𝔸).radius) : summable (λ n, x^n / n!) := +summable_of_summable_norm (norm_exp_series_div_summable_of_mem_ball 𝕂 x hx) + +lemma exp_series_div_has_sum_exp_of_mem_ball [complete_space 𝔸] (x : 𝔸) + (hx : x ∈ emetric.ball (0 : 𝔸) (exp_series 𝕂 𝔸).radius) : has_sum (λ n, x^n / n!) (exp 𝕂 x) := +begin + rw ← exp_series_apply_eq_div' x, + exact exp_series_has_sum_exp_of_mem_ball x hx +end + +variables {𝕂} + lemma exp_neg_of_mem_ball [char_zero 𝕂] [complete_space 𝔸] {x : 𝔸} (hx : x ∈ emetric.ball (0 : 𝔸) (exp_series 𝕂 𝔸).radius) : exp 𝕂 (-x) = (exp 𝕂 x)⁻¹ := @@ -360,17 +364,13 @@ end variables {𝕂 𝔸 𝔹} -section complete_algebra - lemma norm_exp_series_summable (x : 𝔸) : summable (λ n, ∥exp_series 𝕂 𝔸 n (λ _, x)∥) := norm_exp_series_summable_of_mem_ball x ((exp_series_radius_eq_top 𝕂 𝔸).symm ▸ edist_lt_top _ _) lemma norm_exp_series_summable' (x : 𝔸) : summable (λ n, ∥(n!⁻¹ : 𝕂) • x^n∥) := norm_exp_series_summable_of_mem_ball' x ((exp_series_radius_eq_top 𝕂 𝔸).symm ▸ edist_lt_top _ _) -lemma norm_exp_series_field_summable (x : 𝕂) : summable (λ n, ∥x^n / n!∥) := -norm_exp_series_field_summable_of_mem_ball x - ((exp_series_radius_eq_top 𝕂 𝕂).symm ▸ edist_lt_top _ _) +section complete_algebra variables [complete_space 𝔸] @@ -380,18 +380,12 @@ summable_of_summable_norm (norm_exp_series_summable x) lemma exp_series_summable' (x : 𝔸) : summable (λ n, (n!⁻¹ : 𝕂) • x^n) := summable_of_summable_norm (norm_exp_series_summable' x) -lemma exp_series_field_summable (x : 𝕂) : summable (λ n, x^n / n!) := -summable_of_summable_norm (norm_exp_series_field_summable x) - lemma exp_series_has_sum_exp (x : 𝔸) : has_sum (λ n, exp_series 𝕂 𝔸 n (λ _, x)) (exp 𝕂 x) := exp_series_has_sum_exp_of_mem_ball x ((exp_series_radius_eq_top 𝕂 𝔸).symm ▸ edist_lt_top _ _) lemma exp_series_has_sum_exp' (x : 𝔸) : has_sum (λ n, (n!⁻¹ : 𝕂) • x^n) (exp 𝕂 x):= exp_series_has_sum_exp_of_mem_ball' x ((exp_series_radius_eq_top 𝕂 𝔸).symm ▸ edist_lt_top _ _) -lemma exp_series_field_has_sum_exp (x : 𝕂) : has_sum (λ n, x^n / n!) (exp 𝕂 x):= -exp_series_field_has_sum_exp_of_mem_ball x ((exp_series_radius_eq_top 𝕂 𝕂).symm ▸ edist_lt_top _ _) - lemma exp_has_fpower_series_on_ball : has_fpower_series_on_ball (exp 𝕂) (exp_series 𝕂 𝔸) 0 ∞ := exp_series_radius_eq_top 𝕂 𝔸 ▸ @@ -401,8 +395,7 @@ lemma exp_has_fpower_series_at_zero : has_fpower_series_at (exp 𝕂) (exp_series 𝕂 𝔸) 0 := exp_has_fpower_series_on_ball.has_fpower_series_at -lemma exp_continuous : - continuous (exp 𝕂 : 𝔸 → 𝔸) := +lemma exp_continuous : continuous (exp 𝕂 : 𝔸 → 𝔸) := begin rw [continuous_iff_continuous_on_univ, ← metric.eball_top_eq_univ (0 : 𝔸), ← exp_series_radius_eq_top 𝕂 𝔸], @@ -532,8 +525,24 @@ end any_algebra section division_algebra variables {𝕂 𝔸 : Type*} [is_R_or_C 𝕂] [normed_division_ring 𝔸] [normed_algebra 𝕂 𝔸] + +variables (𝕂) + +lemma norm_exp_series_div_summable (x : 𝔸) : summable (λ n, ∥x^n / n!∥) := +norm_exp_series_div_summable_of_mem_ball 𝕂 x + ((exp_series_radius_eq_top 𝕂 𝔸).symm ▸ edist_lt_top _ _) + variables [complete_space 𝔸] +lemma exp_series_div_summable (x : 𝔸) : summable (λ n, x^n / n!) := +summable_of_summable_norm (norm_exp_series_div_summable 𝕂 x) + +lemma exp_series_div_has_sum_exp (x : 𝔸) : has_sum (λ n, x^n / n!) (exp 𝕂 x):= +exp_series_div_has_sum_exp_of_mem_ball 𝕂 x + ((exp_series_radius_eq_top 𝕂 𝔸).symm ▸ edist_lt_top _ _) + +variables {𝕂} + lemma exp_neg (x : 𝔸) : exp 𝕂 (-x) = (exp 𝕂 x)⁻¹ := exp_neg_of_mem_ball $ (exp_series_radius_eq_top 𝕂 𝔸).symm ▸ edist_lt_top _ _ diff --git a/src/analysis/special_functions/exponential.lean b/src/analysis/special_functions/exponential.lean index 4e7e63ad2e68a..7d4ffb38ec4fd 100644 --- a/src/analysis/special_functions/exponential.lean +++ b/src/analysis/special_functions/exponential.lean @@ -207,9 +207,9 @@ section complex lemma complex.exp_eq_exp_ℂ : complex.exp = exp ℂ := begin refine funext (λ x, _), - rw [complex.exp, exp_eq_tsum_field], + rw [complex.exp, exp_eq_tsum_div], exact tendsto_nhds_unique x.exp'.tendsto_limit - (exp_series_field_summable x).has_sum.tendsto_sum_nat + (exp_series_div_summable ℝ x).has_sum.tendsto_sum_nat end end complex @@ -219,7 +219,7 @@ section real lemma real.exp_eq_exp_ℝ : real.exp = exp ℝ := begin refine funext (λ x, _), - rw [real.exp, complex.exp_eq_exp_ℂ, ← exp_ℝ_ℂ_eq_exp_ℂ_ℂ, exp_eq_tsum, exp_eq_tsum_field, + rw [real.exp, complex.exp_eq_exp_ℂ, ← exp_ℝ_ℂ_eq_exp_ℂ_ℂ, exp_eq_tsum, exp_eq_tsum_div, ← re_to_complex, ← re_clm_apply, re_clm.map_tsum (exp_series_summable' (x : ℂ))], refine tsum_congr (λ n, _), rw [re_clm.map_smul, ← complex.of_real_pow, re_clm_apply, re_to_complex, complex.of_real_re, diff --git a/src/analysis/specific_limits/normed.lean b/src/analysis/specific_limits/normed.lean index bb546303e235d..d4262585172aa 100644 --- a/src/analysis/specific_limits/normed.lean +++ b/src/analysis/specific_limits/normed.lean @@ -622,7 +622,7 @@ end ### Factorial -/ -/-- The series `∑' n, x ^ n / n!` is summable of any `x : ℝ`. See also `exp_series_field_summable` +/-- The series `∑' n, x ^ n / n!` is summable of any `x : ℝ`. See also `exp_series_div_summable` for a version that also works in `ℂ`, and `exp_series_summable'` for a version that works in any normed algebra over `ℝ` or `ℂ`. -/ lemma real.summable_pow_div_factorial (x : ℝ) : diff --git a/src/combinatorics/derangements/exponential.lean b/src/combinatorics/derangements/exponential.lean index 3240daf71c8f2..afd91585d11dd 100644 --- a/src/combinatorics/derangements/exponential.lean +++ b/src/combinatorics/derangements/exponential.lean @@ -35,7 +35,7 @@ begin -- there's no specific lemma for ℝ that ∑ x^k/k! sums to exp(x), but it's -- true in more general fields, so use that lemma rw real.exp_eq_exp_ℝ, - exact exp_series_field_has_sum_exp (-1 : ℝ) }, + exact exp_series_div_has_sum_exp ℝ (-1 : ℝ) }, intro n, rw [← int.cast_coe_nat, num_derangements_sum], push_cast, From 8412f1fc16838f696855f6e95575cb8a2f1299e6 Mon Sep 17 00:00:00 2001 From: antoinelab01 Date: Mon, 9 May 2022 02:37:49 +0000 Subject: [PATCH 34/37] feat(representation_theory/invariants): invariants of `lin_hom` are representation morphisms (#14012) --- src/representation_theory/invariants.lean | 33 +++++++++++++++++++++++ 1 file changed, 33 insertions(+) diff --git a/src/representation_theory/invariants.lean b/src/representation_theory/invariants.lean index 53aad38978876..88bbe9353dc05 100644 --- a/src/representation_theory/invariants.lean +++ b/src/representation_theory/invariants.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Antoine Labelle -/ import representation_theory.basic +import representation_theory.Rep /-! # Subspace of invariants a group representation @@ -66,6 +67,8 @@ end group_algebra namespace representation +section invariants + open group_algebra variables {k G V : Type*} [comm_semiring k] [group G] [add_comm_monoid V] [module k V] @@ -111,4 +114,34 @@ begin simp [average_def, map_sum, hv, finset.card_univ, nsmul_eq_smul_cast k _ v, smul_smul], end +end invariants + +namespace lin_hom + +universes u + +open category_theory Action + +variables {k : Type u} [comm_ring k] {G : Group.{u}} + +lemma mem_invariants_iff_comm {X Y : Rep k G} (f : X.V →ₗ[k] Y.V) (g : G) : + (lin_hom X.ρ Y.ρ) g f = f ↔ X.ρ g ≫ f = f ≫ Y.ρ g := +begin + rw [lin_hom_apply, ←ρ_Aut_apply_inv, ←linear_map.comp_assoc, ←Module.comp_def, ←Module.comp_def, + iso.inv_comp_eq, ρ_Aut_apply_hom], exact comm, +end + +/-- The invariants of the representation `lin_hom X.ρ Y.ρ` correspond to the the representation +homomorphisms from `X` to `Y` -/ +@[simps] +def invariants_equiv_Rep_hom (X Y : Rep k G) : (lin_hom X.ρ Y.ρ).invariants ≃ₗ[k] (X ⟶ Y) := +{ to_fun := λ f, ⟨f.val, λ g, (mem_invariants_iff_comm _ g).1 (f.property g)⟩, + map_add' := λ _ _, rfl, + map_smul' := λ _ _, rfl, + inv_fun := λ f, ⟨f.hom, λ g, (mem_invariants_iff_comm _ g).2 (f.comm g)⟩, + left_inv := λ _, by { ext, refl }, + right_inv := λ _, by { ext, refl } } + +end lin_hom + end representation From 77c86ba48f6c444a97fb5e7b7d5fd8edea0b6074 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 9 May 2022 02:37:50 +0000 Subject: [PATCH 35/37] =?UTF-8?q?rename(imo/imo1972=5Fb2=20=E2=86=92=20imo?= =?UTF-8?q?/imo1972=5Fq5):=20Fix=20file=20name=20(#14037)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- archive/imo/{imo1972_b2.lean => imo1972_q5.lean} | 5 ++--- 1 file changed, 2 insertions(+), 3 deletions(-) rename archive/imo/{imo1972_b2.lean => imo1972_q5.lean} (98%) diff --git a/archive/imo/imo1972_b2.lean b/archive/imo/imo1972_q5.lean similarity index 98% rename from archive/imo/imo1972_b2.lean rename to archive/imo/imo1972_q5.lean index 722405f8ccf67..7bd89bacc97d6 100644 --- a/archive/imo/imo1972_b2.lean +++ b/archive/imo/imo1972_q5.lean @@ -8,7 +8,7 @@ import data.real.basic import analysis.normed_space.basic /-! -# IMO 1972 B2 +# IMO 1972 Q5 Problem: `f` and `g` are real-valued functions defined on the real line. For all `x` and `y`, `f(x + y) + f(x - y) = 2f(x)g(y)`. `f` is not identically zero and `|f(x)| ≤ 1` for all `x`. @@ -30,7 +30,6 @@ example (f g : ℝ → ℝ) (y : ℝ) : ∥g(y)∥ ≤ 1 := begin - classical, set S := set.range (λ x, ∥f x∥), -- Introduce `k`, the supremum of `f`. let k : ℝ := Sup (S), @@ -86,7 +85,7 @@ begin end -/-- IMO 1972 B2 +/-- IMO 1972 Q5 Problem: `f` and `g` are real-valued functions defined on the real line. For all `x` and `y`, `f(x + y) + f(x - y) = 2f(x)g(y)`. `f` is not identically zero and `|f(x)| ≤ 1` for all `x`. From bf8db9b77bbd54290360f2eaec587a36d28bebb5 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 9 May 2022 06:03:23 +0000 Subject: [PATCH 36/37] feat(analysis/normed_space/matrix_exponential): lemmas about the matrix exponential (#13520) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This checks off "Matrix Exponential" from the undergrad TODO list, by providing the majority of the "obvious" statements about matrices over a real normed algebra. Combining this PR with what is already in mathlib, we have: * `exp 0 = 1` * `exp (A + B) = exp A * exp B` when `A` and `B` commute * `exp (n • A) = exp A ^ n` * `exp (z • A) = exp A ^ z` * `exp (-A) = (exp A)⁻¹` * `exp (U * D * ↑(U⁻¹)) = U * exp D * ↑(U⁻¹)` * `exp Aᵀ = (exp A)ᵀ` * `exp Aᴴ = (exp A)ᴴ` * `A * exp B = exp B * A` if `A * B = B * A` * `exp (diagonal v) = diagonal (exp v)` * `exp (block_diagonal v) = block_diagonal (exp v)` * `exp (block_diagonal' v) = block_diagonal' (exp v)` Still missing are: * `det (exp A) = exp (trace A)` * `exp A` can be written a weighted sum of powers of `A : matrix n n R` less than `fintype.card n` (an extension of [`matrix.pow_eq_aeval_mod_charpoly`](https://leanprover-community.github.io/mathlib_docs/linear_algebra/matrix/charpoly/coeff.html#matrix.pow_eq_aeval_mod_charpoly)) The proofs in this PR may seem small, but they had a substantial dependency chain: https://github.com/leanprover-community/mathlib/projects/16. It turns out that there's always more missing glue than you think there is. --- docs/undergrad.yaml | 2 +- src/analysis/normed_space/exponential.lean | 8 +- .../normed_space/matrix_exponential.lean | 229 ++++++++++++++++++ 3 files changed, 236 insertions(+), 3 deletions(-) create mode 100644 src/analysis/normed_space/matrix_exponential.lean diff --git a/docs/undergrad.yaml b/docs/undergrad.yaml index 4c527cb63e9bd..3b3c790f4345b 100644 --- a/docs/undergrad.yaml +++ b/docs/undergrad.yaml @@ -71,7 +71,7 @@ Linear algebra: examples: '' Exponential: endomorphism exponential: '' - matrix exponential: 'https://en.wikipedia.org/wiki/Matrix_exponential' + matrix exponential: 'exp' # 2. Group Theory: diff --git a/src/analysis/normed_space/exponential.lean b/src/analysis/normed_space/exponential.lean index 6fab164c2a7b2..874a65bcbcad7 100644 --- a/src/analysis/normed_space/exponential.lean +++ b/src/analysis/normed_space/exponential.lean @@ -76,10 +76,14 @@ def exp_series : formal_multilinear_series 𝕂 𝔸 𝔸 := variables {𝔸} /-- `exp 𝕂 : 𝔸 → 𝔸` is the exponential map determined by the action of `𝕂` on `𝔸`. -It is defined as the sum of the `formal_multilinear_series` `exp_series 𝕂 𝔸`. -/ +It is defined as the sum of the `formal_multilinear_series` `exp_series 𝕂 𝔸`. + +Note that when `𝔸 = matrix n n 𝕂`, this is the **Matrix Exponential**; see +[`analysis.normed_space.matrix_exponential`](../matrix_exponential) for lemmas specific to that +case. -/ noncomputable def exp (x : 𝔸) : 𝔸 := (exp_series 𝕂 𝔸).sum x -variables {𝕂 𝔸} +variables {𝕂} lemma exp_series_apply_eq (x : 𝔸) (n : ℕ) : exp_series 𝕂 𝔸 n (λ _, x) = (n!⁻¹ : 𝕂) • x^n := by simp [exp_series] diff --git a/src/analysis/normed_space/matrix_exponential.lean b/src/analysis/normed_space/matrix_exponential.lean new file mode 100644 index 0000000000000..9426abbcf0608 --- /dev/null +++ b/src/analysis/normed_space/matrix_exponential.lean @@ -0,0 +1,229 @@ +/- +Copyright (c) 2022 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ + +import analysis.normed_space.exponential +import analysis.matrix +import linear_algebra.matrix.zpow +import topology.uniform_space.matrix + +/-! +# Lemmas about the matrix exponential + +In this file, we provide results about `exp` on `matrix`s over a topological or normed algebra. +Note that generic results over all topological spaces such as `exp_zero` can be used on matrices +without issue, so are not repeated here. The topological results specific to matrices are: + +* `matrix.exp_transpose` +* `matrix.exp_conj_transpose` +* `matrix.exp_diagonal` +* `matrix.exp_block_diagonal` +* `matrix.exp_block_diagonal'` + +Lemmas like `exp_add_of_commute` require a canonical norm on the type; while there are multiple +sensible choices for the norm of a `matrix` (`matrix.normed_group`, `matrix.frobenius_normed_group`, +`matrix.linfty_op_normed_group`), none of them are canonical. In an application where a particular +norm is chosen using `local attribute [instance]`, then the usual lemmas about `exp` are fine. When +choosing a norm is undesirable, the results in this file can be used. + +In this file, we copy across the lemmas about `exp`, but hide the requirement for a norm inside the +proof. + +* `matrix.exp_add_of_commute` +* `matrix.exp_sum_of_commute` +* `matrix.exp_nsmul` +* `matrix.is_unit_exp` +* `matrix.exp_units_conj` +* `matrix.exp_units_conj'` + +Additionally, we prove some results about `matrix.has_inv` and `matrix.div_inv_monoid`, as the +results for general rings are instead stated about `ring.inverse`: + +* `matrix.exp_neg` +* `matrix.exp_zsmul` +* `matrix.exp_conj` +* `matrix.exp_conj'` + +## Implementation notes + +This file runs into some sharp edges on typeclass search in lean 3, especially regarding pi types. +To work around this, we copy a handful of instances for when lean can't find them by itself. +Hopefully we will be able to remove these in Lean 4. + +## TODO + +* Show that `matrix.det (exp 𝕂 A) = exp 𝕂 (matrix.trace A)` + +## References + +* https://en.wikipedia.org/wiki/Matrix_exponential +-/ +open_locale matrix big_operators + +section hacks_for_pi_instance_search + +/-- A special case of `pi.topological_ring` for when `R` is not dependently typed. -/ +instance function.topological_ring (I : Type*) (R : Type*) + [non_unital_ring R] [topological_space R] [topological_ring R] : + topological_ring (I → R) := +pi.topological_ring + +/-- A special case of `function.algebra` for when A is a `ring` not a `semiring` -/ +instance function.algebra_ring (I : Type*) {R : Type*} (A : Type*) [comm_semiring R] + [ring A] [algebra R A] : algebra R (I → A) := +pi.algebra _ _ + +/-- A special case of `pi.algebra` for when `f = λ i, matrix (m i) (m i) A`. -/ +instance pi.matrix_algebra (I R A : Type*) (m : I → Type*) + [comm_semiring R] [semiring A] [algebra R A] + [Π i, fintype (m i)] [Π i, decidable_eq (m i)] : + algebra R (Π i, matrix (m i) (m i) A) := +@pi.algebra I R (λ i, matrix (m i) (m i) A) _ _ (λ i, matrix.algebra) + +/-- A special case of `pi.topological_ring` for when `f = λ i, matrix (m i) (m i) A`. -/ +instance pi.matrix_topological_ring (I A : Type*) (m : I → Type*) + [ring A] [topological_space A] [topological_ring A] + [Π i, fintype (m i)] : + topological_ring (Π i, matrix (m i) (m i) A) := +@pi.topological_ring _ (λ i, matrix (m i) (m i) A) _ _ (λ i, matrix.topological_ring) + +end hacks_for_pi_instance_search + +variables (𝕂 : Type*) {m n p : Type*} {n' : m → Type*} {𝔸 : Type*} + +namespace matrix + +section topological + +section ring +variables [fintype m] [decidable_eq m] [fintype n] [decidable_eq n] + [Π i, fintype (n' i)] [Π i, decidable_eq (n' i)] + [field 𝕂] [ring 𝔸] [topological_space 𝔸] [topological_ring 𝔸] [algebra 𝕂 𝔸] [t2_space 𝔸] + +lemma exp_diagonal (v : m → 𝔸) : exp 𝕂 (diagonal v) = diagonal (exp 𝕂 v) := +by simp_rw [exp_eq_tsum, diagonal_pow, ←diagonal_smul, ←diagonal_tsum] + +lemma exp_block_diagonal (v : m → matrix n n 𝔸) : + exp 𝕂 (block_diagonal v) = block_diagonal (exp 𝕂 v) := +by simp_rw [exp_eq_tsum, ←block_diagonal_pow, ←block_diagonal_smul, ←block_diagonal_tsum] + +lemma exp_block_diagonal' (v : Π i, matrix (n' i) (n' i) 𝔸) : + exp 𝕂 (block_diagonal' v) = block_diagonal' (exp 𝕂 v) := +by simp_rw [exp_eq_tsum, ←block_diagonal'_pow, ←block_diagonal'_smul, ←block_diagonal'_tsum] + +lemma exp_conj_transpose [star_ring 𝔸] [has_continuous_star 𝔸] (A : matrix m m 𝔸) : + exp 𝕂 Aᴴ = (exp 𝕂 A)ᴴ := +(star_exp A).symm + +end ring + +section comm_ring +variables [fintype m] [decidable_eq m] [field 𝕂] + [comm_ring 𝔸] [topological_space 𝔸] [topological_ring 𝔸] [algebra 𝕂 𝔸] [t2_space 𝔸] + +lemma exp_transpose (A : matrix m m 𝔸) : exp 𝕂 Aᵀ = (exp 𝕂 A)ᵀ := +by simp_rw [exp_eq_tsum, transpose_tsum, transpose_smul, transpose_pow] + +end comm_ring + +end topological + +section normed + +variables [is_R_or_C 𝕂] + [fintype m] [decidable_eq m] + [fintype n] [decidable_eq n] + [Π i, fintype (n' i)] [Π i, decidable_eq (n' i)] + [normed_ring 𝔸] [normed_algebra 𝕂 𝔸] [complete_space 𝔸] + +lemma exp_add_of_commute (A B : matrix m m 𝔸) (h : commute A B) : + exp 𝕂 (A + B) = exp 𝕂 A ⬝ exp 𝕂 B := +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, + letI : normed_algebra 𝕂 (matrix m m 𝔸) := matrix.linfty_op_normed_algebra, + exact exp_add_of_commute h, +end + +lemma exp_sum_of_commute {ι} (s : finset ι) (f : ι → matrix m m 𝔸) + (h : ∀ (i ∈ s) (j ∈ s), 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 𝕂) := +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, + letI : normed_algebra 𝕂 (matrix m m 𝔸) := matrix.linfty_op_normed_algebra, + exact exp_sum_of_commute s f h, +end + +lemma exp_nsmul (n : ℕ) (A : matrix m m 𝔸) : + exp 𝕂 (n • A) = exp 𝕂 A ^ n := +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, + letI : normed_algebra 𝕂 (matrix m m 𝔸) := matrix.linfty_op_normed_algebra, + exact exp_nsmul n A, +end + +lemma is_unit_exp (A : matrix m m 𝔸) : is_unit (exp 𝕂 A) := +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, + letI : normed_algebra 𝕂 (matrix m m 𝔸) := matrix.linfty_op_normed_algebra, + exact is_unit_exp _ A, +end + +lemma exp_units_conj (U : (matrix m m 𝔸)ˣ) (A : matrix m m 𝔸) : + exp 𝕂 (↑U ⬝ A ⬝ ↑(U⁻¹) : matrix m m 𝔸) = ↑U ⬝ exp 𝕂 A ⬝ ↑(U⁻¹) := +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, + letI : normed_algebra 𝕂 (matrix m m 𝔸) := matrix.linfty_op_normed_algebra, + exact exp_units_conj _ U A, +end + +lemma exp_units_conj' (U : (matrix m m 𝔸)ˣ) (A : matrix m m 𝔸) : + exp 𝕂 (↑(U⁻¹) ⬝ A ⬝ U : matrix m m 𝔸) = ↑(U⁻¹) ⬝ exp 𝕂 A ⬝ U := +exp_units_conj 𝕂 U⁻¹ A + +end normed + +section normed_comm + +variables [is_R_or_C 𝕂] + [fintype m] [decidable_eq m] + [fintype n] [decidable_eq n] + [Π i, fintype (n' i)] [Π i, decidable_eq (n' i)] + [normed_comm_ring 𝔸] [normed_algebra 𝕂 𝔸] [complete_space 𝔸] + +lemma exp_neg (A : matrix m m 𝔸) : exp 𝕂 (-A) = (exp 𝕂 A)⁻¹ := +begin + rw nonsing_inv_eq_ring_inverse, + letI : semi_normed_ring (matrix m m 𝔸) := matrix.linfty_op_semi_normed_ring, + letI : normed_ring (matrix m m 𝔸) := matrix.linfty_op_normed_ring, + letI : normed_algebra 𝕂 (matrix m m 𝔸) := matrix.linfty_op_normed_algebra, + exact (ring.inverse_exp _ A).symm, +end + +lemma exp_zsmul (z : ℤ) (A : matrix m m 𝔸) : exp 𝕂 (z • A) = exp 𝕂 A ^ z := +begin + obtain ⟨n, rfl | rfl⟩ := z.eq_coe_or_neg, + { rw [zpow_coe_nat, coe_nat_zsmul, exp_nsmul] }, + { have : is_unit (exp 𝕂 A).det := (matrix.is_unit_iff_is_unit_det _).mp (is_unit_exp _ _), + rw [matrix.zpow_neg this, zpow_coe_nat, neg_smul, + exp_neg, coe_nat_zsmul, exp_nsmul] }, +end + +lemma exp_conj (U : matrix m m 𝔸) (A : matrix m m 𝔸) (hy : is_unit U) : + exp 𝕂 (U ⬝ A ⬝ U⁻¹) = U ⬝ exp 𝕂 A ⬝ U⁻¹ := +let ⟨u, hu⟩ := hy in hu ▸ by simpa only [matrix.coe_units_inv] using exp_units_conj 𝕂 u A + +lemma exp_conj' (U : matrix m m 𝔸) (A : matrix m m 𝔸) (hy : is_unit U) : + exp 𝕂 (U⁻¹ ⬝ A ⬝ U) = U⁻¹ ⬝ exp 𝕂 A ⬝ U := +let ⟨u, hu⟩ := hy in hu ▸ by simpa only [matrix.coe_units_inv] using exp_units_conj' 𝕂 u A + +end normed_comm + +end matrix From b38aee47c8131c8ff6802b90de10d61bd4dc3987 Mon Sep 17 00:00:00 2001 From: loefflerd Date: Mon, 9 May 2022 08:31:21 +0000 Subject: [PATCH 37/37] feat(analysis/special_functions): differentiability of Gamma function (#13000) Third instalment of my Gamma-function project (following #12917 and #13156). This PR adds the proof that the Gamma function is complex-analytic, away from the poles at non-positive integers. --- src/analysis/special_functions/gamma.lean | 202 ++++++++++++++++-- src/analysis/special_functions/log/basic.lean | 13 ++ src/analysis/special_functions/pow.lean | 9 + 3 files changed, 212 insertions(+), 12 deletions(-) diff --git a/src/analysis/special_functions/gamma.lean b/src/analysis/special_functions/gamma.lean index f94a333f6ad96..9ed8844250991 100644 --- a/src/analysis/special_functions/gamma.lean +++ b/src/analysis/special_functions/gamma.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: David Loeffler -/ import measure_theory.integral.exp_decay +import analysis.calculus.parametric_integral /-! # The Gamma function @@ -18,15 +19,13 @@ We show that this integral satisfies `Γ(1) = 1` and `Γ(s + 1) = s * Γ(s)`; he `Γ(s)` for all `s` as the unique function satisfying this recurrence and agreeing with Euler's integral in the convergence range. -TODO: Holomorpy in `s` (away from the poles at `-n : n ∈ ℕ`) will be added in a future PR. - ## Tags Gamma -/ noncomputable theory -open filter interval_integral set real measure_theory +open filter interval_integral set real measure_theory asymptotics open_locale topological_space lemma integral_exp_neg_Ioi : ∫ (x : ℝ) in Ioi 0, exp (-x) = 1 := @@ -38,11 +37,11 @@ end namespace real -/-- Asymptotic bound for the Γ function integrand. -/ -lemma Gamma_integrand_is_O (s : ℝ) : asymptotics.is_O (λ x:ℝ, exp (-x) * x ^ s) +/-- Asymptotic bound for the `Γ` function integrand. -/ +lemma Gamma_integrand_is_O (s : ℝ) : is_O (λ x:ℝ, exp (-x) * x ^ s) (λ x:ℝ, exp (-(1/2) * x)) at_top := begin - refine asymptotics.is_o.is_O (asymptotics.is_o_of_tendsto _ _), + refine is_o.is_O (is_o_of_tendsto _ _), { intros x hx, exfalso, exact (exp_pos (-(1 / 2) * x)).ne' hx }, have : (λ (x:ℝ), exp (-x) * x ^ s / exp (-(1 / 2) * x)) = (λ (x:ℝ), exp ((1 / 2) * x) / x ^ s )⁻¹, { ext1 x, @@ -59,7 +58,7 @@ end See `Gamma_integral_convergent` for a proof of the convergence of the integral for `1 ≤ s`. -/ def Gamma_integral (s : ℝ) : ℝ := ∫ x in Ioi (0:ℝ), exp (-x) * x ^ (s - 1) -/-- The integral defining the Γ function converges for real `s` with `1 ≤ s`. +/-- The integral defining the `Γ` function converges for real `s` with `1 ≤ s`. This is not optimal, but the optimal bound (convergence for `0 < s`) is hard to establish with the results currently in the library. -/ @@ -82,7 +81,7 @@ make a choice between ↑(real.exp (-x)), complex.exp (↑(-x)), and complex.exp equal but not definitionally so. We use the first of these throughout. -/ -/-- The integral defining the Γ function converges for complex `s` with `1 ≤ re s`. +/-- The integral defining the `Γ` function converges for complex `s` with `1 ≤ re s`. This is proved by reduction to the real case. The bound is not optimal, but the optimal bound (convergence for `0 < re s`) is hard to establish with the results currently in the library. -/ @@ -140,7 +139,7 @@ namespace complex section Gamma_recurrence -/-- The indefinite version of the Γ function, Γ(s, X) = ∫ x ∈ 0..X, exp(-x) x ^ (s - 1). -/ +/-- The indefinite version of the `Γ` function, `Γ(s, X) = ∫ x ∈ 0..X, exp(-x) x ^ (s - 1)`. -/ def partial_Gamma (s : ℂ) (X : ℝ) : ℂ := ∫ x in 0..X, (-x).exp * x ^ (s - 1) lemma tendsto_partial_Gamma {s : ℂ} (hs: 1 ≤ s.re) : @@ -186,7 +185,7 @@ begin simpa using mul_le_mul i1 i2 (abs_nonneg (↑x ^ (s - 1))) zero_le_one, end -/-- The recurrence relation for the indefinite version of the Γ function. -/ +/-- The recurrence relation for the indefinite version of the `Γ` function. -/ lemma partial_Gamma_add_one {s : ℂ} (hs: 1 ≤ s.re) {X : ℝ} (hX : 0 ≤ X) : partial_Gamma (s + 1) X = s * partial_Gamma s X - (-X).exp * X ^ s := begin @@ -226,7 +225,7 @@ begin { contrapose! hs, rw [hs, zero_re], exact zero_lt_one,} end -/-- The recurrence relation for the Γ integral. -/ +/-- The recurrence relation for the `Γ` integral. -/ theorem Gamma_integral_add_one {s : ℂ} (hs: 1 ≤ s.re) : Gamma_integral (s + 1) = s * Gamma_integral s := begin @@ -310,7 +309,7 @@ begin rw [←nat.add_sub_of_le (nat.ceil_le.mpr h1), u (n - ⌈ 1 - s.re ⌉₊)], end -/-- The recurrence relation for the Γ function. -/ +/-- The recurrence relation for the `Γ` function. -/ theorem Gamma_add_one (s : ℂ) (h2 : s ≠ 0) : Gamma (s+1) = s * Gamma s := begin let n := ⌈ 1 - s.re ⌉₊, @@ -337,3 +336,182 @@ end end Gamma_def end complex + +/-! Now check that the `Γ` function is differentiable, wherever this makes sense. -/ + +section Gamma_has_deriv + +/-- Integrand for the derivative of the `Γ` function -/ +def dGamma_integrand (s : ℂ) (x : ℝ) : ℂ := exp (-x) * log x * x ^ (s - 1) + +/-- Integrand for the absolute value of the derivative of the `Γ` function -/ +def dGamma_integrand_real (s x : ℝ) : ℝ := |exp (-x) * log x * x ^ (s - 1)| + +lemma dGamma_integrand_is_O_at_top (s : ℝ) : is_O (λ x:ℝ, exp (-x) * log x * x ^ (s - 1)) + (λ x:ℝ, exp (-(1/2) * x)) at_top := +begin + refine is_o.is_O (is_o_of_tendsto _ _), + { intros x hx, exfalso, exact (-(1/2) * x).exp_pos.ne' hx, }, + have : eventually_eq at_top (λ (x : ℝ), exp (-x) * log x * x ^ (s - 1) / exp (-(1 / 2) * x)) + (λ (x : ℝ), (λ z:ℝ, exp (1 / 2 * z) / z ^ s) x * (λ z:ℝ, z / log z) x)⁻¹, + { refine eventually_of_mem (Ioi_mem_at_top 1) _, + intros x hx, dsimp, + replace hx := lt_trans zero_lt_one (mem_Ioi.mp hx), + rw [real.exp_neg, neg_mul, real.exp_neg, rpow_sub hx], + have : exp x = exp(x/2) * exp(x/2) := by { rw ←real.exp_add, simp, }, rw this, + field_simp [hx.ne', exp_ne_zero (x/2)], ring, }, + refine tendsto.congr' this.symm (tendsto.inv_tendsto_at_top _), + apply tendsto.at_top_mul_at_top (tendsto_exp_mul_div_rpow_at_top s (1/2) one_half_pos), + refine tendsto.congr' _ ((tendsto_exp_div_pow_at_top 1).comp tendsto_log_at_top), + apply eventually_eq_of_mem (Ioi_mem_at_top (0:ℝ)), + intros x hx, simp [exp_log hx], +end + +/-- Absolute convergence of the integral which will give the derivative of the `Γ` function on +`1 < re s`. -/ +lemma dGamma_integral_abs_convergent (s : ℝ) (hs : 1 < s) : + integrable_on (λ x:ℝ, ∥exp (-x) * log x * x ^ (s-1)∥) (Ioi 0) := +begin + have : Ioi (0:ℝ) = Ioc 0 1 ∪ Ioi 1 := by simp, + rw [this, integrable_on_union], + refine ⟨⟨_, _⟩, _⟩, + { refine continuous_on.ae_strongly_measurable (continuous_on.mul _ _).norm measurable_set_Ioc, + { refine (continuous_exp.comp continuous_neg).continuous_on.mul (continuous_on_log.mono _), + simp, }, + { apply continuous_on_id.rpow_const, intros x hx, right, linarith }, }, + { apply has_finite_integral_of_bounded, + swap, { exact 1 / (s - 1), }, + refine (ae_restrict_iff' measurable_set_Ioc).mpr (ae_of_all _ (λ x hx, _)), + rw [norm_norm, norm_eq_abs, mul_assoc, abs_mul], + have : 1/(s-1) = 1 * (1 / (s-1)) := by ring, rw this, + refine mul_le_mul _ _ (by apply abs_nonneg) zero_le_one, + { rw [abs_of_pos (exp_pos(-x)), exp_le_one_iff, neg_le, neg_zero], exact hx.1.le }, + { apply le_of_lt, refine abs_log_mul_self_rpow_lt x (s-1) hx.1 hx.2 (by linarith), }, }, + { have := is_O.norm_left (dGamma_integrand_is_O_at_top s), + refine integrable_of_is_O_exp_neg one_half_pos (continuous_on.mul _ _).norm this, + { refine (continuous_exp.comp continuous_neg).continuous_on.mul (continuous_on_log.mono _), + simp, }, + { apply continuous_at.continuous_on (λ x hx, _), + apply continuous_at_id.rpow continuous_at_const, + dsimp, right, linarith, }, } +end + +/-- A uniform bound for the `s`-derivative of the `Γ` integrand for `s` in vertical strips. -/ +lemma loc_unif_bound_dGamma_integrand {t : ℂ} {s1 s2 x : ℝ} (ht1 : s1 ≤ t.re) + (ht2: t.re ≤ s2) (hx : 0 < x) : + ∥dGamma_integrand t x∥ ≤ dGamma_integrand_real s1 x + dGamma_integrand_real s2 x := +begin + rcases le_or_lt 1 x with h|h, + { suffices: ∥dGamma_integrand t x∥ ≤ dGamma_integrand_real s2 x, -- case 1 ≤ x + { have: 0 ≤ dGamma_integrand_real s1 x := by apply abs_nonneg, linarith, }, + rw [dGamma_integrand, dGamma_integrand_real, complex.norm_eq_abs, complex.abs_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, + refine le_trans _ (le_abs_self _), + apply rpow_le_rpow_of_exponent_le h, + rw [complex.sub_re, complex.one_re], linarith, }, + { suffices: ∥dGamma_integrand t x∥ ≤ dGamma_integrand_real s1 x, + { have : 0 ≤ dGamma_integrand_real s2 x := by apply abs_nonneg, linarith, }, + rw [dGamma_integrand, dGamma_integrand_real, complex.norm_eq_abs, complex.abs_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, + refine le_trans _ (le_abs_self _), + apply rpow_le_rpow_of_exponent_ge hx h.le, + rw [complex.sub_re, complex.one_re], linarith, }, +end + +namespace complex + +/-- The derivative of the `Γ` integral, at any `s ∈ ℂ` with `1 < re s`, is given by the integral +of `exp (-x) * log x * x ^ (s - 1)` over `[0, ∞)`. -/ +theorem has_deriv_at_Gamma_integral {s : ℂ} (hs : 1 < s.re) : + (integrable_on (λ x, real.exp (-x) * real.log x * x ^ (s - 1) : ℝ → ℂ) (Ioi 0) volume) ∧ + (has_deriv_at Gamma_integral (∫ x:ℝ in Ioi 0, real.exp (-x) * real.log x * x ^ (s - 1)) s) := +begin + let ε := (s.re - 1) / 2, + let μ := volume.restrict (Ioi (0:ℝ)), + let bound := (λ x:ℝ, dGamma_integrand_real (s.re - ε) x + dGamma_integrand_real (s.re + ε) x), + have cont : ∀ (t : ℂ), continuous_on (λ x, real.exp (-x) * x ^ (t - 1) : ℝ → ℂ) (Ioi 0), + { intro t, apply (continuous_of_real.comp continuous_neg.exp).continuous_on.mul, + apply continuous_at.continuous_on, intros x hx, + refine (continuous_at_cpow_const _).comp continuous_of_real.continuous_at, + exact or.inl hx, }, + have eps_pos: 0 < ε := by { refine div_pos _ zero_lt_two, linarith }, + have hF_meas : ∀ᶠ (t : ℂ) in 𝓝 s, + ae_strongly_measurable (λ x, real.exp(-x) * x ^ (t - 1) : ℝ → ℂ) μ, + { apply eventually_of_forall, intro t, + exact (cont t).ae_strongly_measurable measurable_set_Ioi, }, + have hF'_meas : ae_strongly_measurable (dGamma_integrand s) μ, + { refine continuous_on.ae_strongly_measurable _ measurable_set_Ioi, + have : dGamma_integrand s = (λ x, real.exp (-x) * x ^ (s - 1) * real.log x : ℝ → ℂ), + { ext1, simp only [dGamma_integrand], ring }, + rw this, + refine continuous_on.mul (cont s) (continuous_at.continuous_on _), + exact λ x hx, continuous_of_real.continuous_at.comp (continuous_at_log (mem_Ioi.mp hx).ne'), }, + have h_bound : ∀ᵐ (x : ℝ) ∂μ, ∀ (t : ℂ), t ∈ metric.ball s ε → ∥ dGamma_integrand t x ∥ ≤ bound x, + { refine (ae_restrict_iff' measurable_set_Ioi).mpr (ae_of_all _ (λ x hx, _)), + intros t ht, + rw [metric.mem_ball, complex.dist_eq] at ht, + replace ht := lt_of_le_of_lt (complex.abs_re_le_abs $ t - s ) ht, + rw [complex.sub_re, @abs_sub_lt_iff ℝ _ t.re s.re ((s.re - 1) / 2) ] at ht, + refine loc_unif_bound_dGamma_integrand _ _ hx, + all_goals { simp only [ε], linarith } }, + have bound_integrable : integrable bound μ, + { apply integrable.add, + { refine dGamma_integral_abs_convergent (s.re - ε) _, + field_simp, rw one_lt_div, + { linarith }, { exact zero_lt_two }, }, + { refine dGamma_integral_abs_convergent (s.re + ε) _, linarith, }, }, + have h_diff : ∀ᵐ (x : ℝ) ∂μ, ∀ (t : ℂ), t ∈ metric.ball s ε + → has_deriv_at (λ u, real.exp (-x) * x ^ (u - 1) : ℂ → ℂ) (dGamma_integrand t x) t, + { refine (ae_restrict_iff' measurable_set_Ioi).mpr (ae_of_all _ (λ x hx, _)), + intros t ht, rw mem_Ioi at hx, + simp only [dGamma_integrand], + rw mul_assoc, + apply has_deriv_at.const_mul, + rw [of_real_log hx.le, mul_comm], + have := ((has_deriv_at_id t).sub_const 1).const_cpow (or.inl (of_real_ne_zero.mpr hx.ne')), + rwa mul_one at this }, + exact (has_deriv_at_integral_of_dominated_loc_of_deriv_le eps_pos hF_meas + (Gamma_integral_convergent hs.le) hF'_meas h_bound bound_integrable h_diff), +end + +lemma differentiable_at_Gamma_aux (s : ℂ) (n : ℕ) (h1 : (1 - s.re) < n ) (h2 : ∀ m:ℕ, s + m ≠ 0) : + differentiable_at ℂ (Gamma_aux n) s := +begin + induction n with n hn generalizing s, + { refine (has_deriv_at_Gamma_integral _).2.differentiable_at, + rw nat.cast_zero at h1, linarith }, + { dsimp only [Gamma_aux], + specialize hn (s + 1), + have a : 1 - (s + 1).re < ↑n, + { rw nat.cast_succ at h1, rw [complex.add_re, complex.one_re], linarith }, + have b : ∀ m:ℕ, s + 1 + m ≠ 0, + { intro m, have := h2 (1 + m), rwa [nat.cast_add, nat.cast_one, ←add_assoc] at this }, + refine differentiable_at.div (differentiable_at.comp _ (hn a b) _) _ _, + simp, simp, simpa using h2 0 } +end + +theorem differentiable_at_Gamma (s : ℂ) (hs : ∀ m:ℕ, s + m ≠ 0) : differentiable_at ℂ Gamma s := +begin + let n := ⌊1 - s.re⌋₊ + 1, + have hn : 1 - s.re < n := nat.lt_floor_add_one (1 - s.re), + apply (differentiable_at_Gamma_aux s n hn hs).congr_of_eventually_eq, + let S := { t : ℂ | 1 - t.re < n }, + have : S ∈ 𝓝 s, + { rw mem_nhds_iff, use S, + refine ⟨by refl, _, hn⟩, + have: S = re⁻¹' Ioi (1 - n : ℝ), + { ext, rw [preimage,Ioi, mem_set_of_eq, mem_set_of_eq, mem_set_of_eq], exact sub_lt }, + rw this, + refine continuous.is_open_preimage continuous_re _ is_open_Ioi, }, + apply eventually_eq_of_mem this, + intros t ht, rw mem_set_of_eq at ht, + apply Gamma_eq_Gamma_aux, exact ht.le, +end + +end complex + +end Gamma_has_deriv diff --git a/src/analysis/special_functions/log/basic.lean b/src/analysis/special_functions/log/basic.lean index 91f8e7b559953..9dd6c92a4169e 100644 --- a/src/analysis/special_functions/log/basic.lean +++ b/src/analysis/special_functions/log/basic.lean @@ -196,6 +196,19 @@ begin rw exp_log hx, end +/-- Bound for `|log x * x|` in the interval `(0, 1]`. -/ +lemma abs_log_mul_self_lt (x: ℝ) (h1 : 0 < x) (h2 : x ≤ 1) : |log x * x| < 1 := +begin + have : 0 < 1/x := by simpa only [one_div, inv_pos] using h1, + replace := log_le_sub_one_of_pos this, + replace : log (1 / x) < 1/x := by linarith, + rw [log_div one_ne_zero h1.ne', log_one, zero_sub, lt_div_iff h1] at this, + have aux : 0 ≤ -log x * x, + { refine mul_nonneg _ h1.le, rw ←log_inv, apply log_nonneg, + rw [←(le_inv h1 zero_lt_one), inv_one], exact h2, }, + rw [←(abs_of_nonneg aux), neg_mul, abs_neg] at this, exact this, +end + /-- The real logarithm function tends to `+∞` at `+∞`. -/ lemma tendsto_log_at_top : tendsto log at_top at_top := tendsto_comp_exp_at_top.1 $ by simpa only [log_exp] using tendsto_id diff --git a/src/analysis/special_functions/pow.lean b/src/analysis/special_functions/pow.lean index f24b1db182652..9823375649641 100644 --- a/src/analysis/special_functions/pow.lean +++ b/src/analysis/special_functions/pow.lean @@ -726,6 +726,15 @@ end lemma rpow_le_one_iff_of_pos (hx : 0 < x) : x ^ y ≤ 1 ↔ 1 ≤ x ∧ y ≤ 0 ∨ x ≤ 1 ∧ 0 ≤ y := by rw [rpow_def_of_pos hx, exp_le_one_iff, mul_nonpos_iff, log_nonneg_iff hx, log_nonpos_iff hx] +/-- Bound for `|log x * x ^ t|` in the interval `(0, 1]`, for positive real `t`. -/ +lemma abs_log_mul_self_rpow_lt (x t : ℝ) (h1 : 0 < x) (h2 : x ≤ 1) (ht : 0 < t) : + |log x * x ^ t| < 1 / t := +begin + rw lt_div_iff ht, + have := abs_log_mul_self_lt (x ^ t) (rpow_pos_of_pos h1 t) (rpow_le_one h1.le h2 ht.le), + rwa [log_rpow h1, mul_assoc, abs_mul, abs_of_pos ht, mul_comm] at this +end + lemma pow_nat_rpow_nat_inv {x : ℝ} (hx : 0 ≤ x) {n : ℕ} (hn : 0 < n) : (x ^ n) ^ (n⁻¹ : ℝ) = x := have hn0 : (n : ℝ) ≠ 0, by simpa [pos_iff_ne_zero] using hn,