From fe445767f40c5f55ba36722675d6fcc3fd1f49e4 Mon Sep 17 00:00:00 2001 From: Kexing Ying Date: Mon, 2 May 2022 09:54:04 +0000 Subject: [PATCH 001/153] feat(probability/martingale): the optional stopping theorem (#13630) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We prove the optional stopping theorem (also known as the fair game theorem). This is number 62 on Freek 100 theorems. Co-authored-by: RemyDegenne Co-authored-by: Rémy Degenne --- docs/100.yaml | 2 + src/algebra/indicator_function.lean | 13 ++++ src/measure_theory/integral/bochner.lean | 18 ++++- src/measure_theory/integral/set_integral.lean | 7 ++ src/measure_theory/measure/measure_space.lean | 9 +++ src/probability/martingale.lean | 57 ++++++++++++++ src/probability/stopping.lean | 76 +++++++++++++++++++ 7 files changed, 181 insertions(+), 1 deletion(-) diff --git a/docs/100.yaml b/docs/100.yaml index d8777ff14830b..4003c7c42a912 100644 --- a/docs/100.yaml +++ b/docs/100.yaml @@ -211,6 +211,8 @@ title : Theorem of Ceva 62: title : Fair Games Theorem + decl : measure_theory.submartingale_iff_expected_stopped_value_mono + author : Kexing Ying 63: title : Cantor’s Theorem decl : cardinal.cantor diff --git a/src/algebra/indicator_function.lean b/src/algebra/indicator_function.lean index ce051b09228c5..25bde364046b8 100644 --- a/src/algebra/indicator_function.lean +++ b/src/algebra/indicator_function.lean @@ -249,6 +249,19 @@ begin rw [pi.mul_apply, this, one_mul] end +@[to_additive] lemma mul_indicator_mul_compl_eq_piecewise + [decidable_pred (∈ s)] (f g : α → M) : + s.mul_indicator f * sᶜ.mul_indicator g = s.piecewise f g := +begin + ext x, + by_cases h : x ∈ s, + { rw [piecewise_eq_of_mem _ _ _ h, pi.mul_apply, set.mul_indicator_of_mem h, + set.mul_indicator_of_not_mem (set.not_mem_compl_iff.2 h), mul_one] }, + { rw [piecewise_eq_of_not_mem _ _ _ h, pi.mul_apply, set.mul_indicator_of_not_mem h, + set.mul_indicator_of_mem (set.mem_compl h), one_mul] }, +end + + /-- `set.mul_indicator` as a `monoid_hom`. -/ @[to_additive "`set.indicator` as an `add_monoid_hom`."] def mul_indicator_hom {α} (M) [mul_one_class M] (s : set α) : (α → M) →* (α → M) := diff --git a/src/measure_theory/integral/bochner.lean b/src/measure_theory/integral/bochner.lean index 21271e34c0b30..5c7bebc92e623 100644 --- a/src/measure_theory/integral/bochner.lean +++ b/src/measure_theory/integral/bochner.lean @@ -1456,7 +1456,7 @@ lemma ae_eq_trim_of_strongly_measurable f =ᵐ[μ.trim hm] g := begin rwa [eventually_eq, ae_iff, trim_measurable_set_eq hm _], - exact measurable_set.compl (hf.measurable_set_eq_fun hg) + exact (hf.measurable_set_eq_fun hg).compl end lemma ae_eq_trim_iff [topological_space γ] [metrizable_space γ] @@ -1464,6 +1464,22 @@ lemma ae_eq_trim_iff [topological_space γ] [metrizable_space γ] f =ᵐ[μ.trim hm] g ↔ f =ᵐ[μ] g := ⟨ae_eq_of_ae_eq_trim, ae_eq_trim_of_strongly_measurable hm hf hg⟩ +lemma ae_le_trim_of_strongly_measurable + [linear_order γ] [topological_space γ] [order_closed_topology γ] [metrizable_space γ] + (hm : m ≤ m0) {f g : β → γ} (hf : strongly_measurable[m] f) (hg : strongly_measurable[m] g) + (hfg : f ≤ᵐ[μ] g) : + f ≤ᵐ[μ.trim hm] g := +begin + rwa [eventually_le, ae_iff, trim_measurable_set_eq hm _], + exact (hf.measurable_set_le hg).compl, +end + +lemma ae_le_trim_iff + [linear_order γ] [topological_space γ] [order_closed_topology γ] [metrizable_space γ] + (hm : m ≤ m0) {f g : β → γ} (hf : strongly_measurable[m] f) (hg : strongly_measurable[m] g) : + f ≤ᵐ[μ.trim hm] g ↔ f ≤ᵐ[μ] g := +⟨ae_le_of_ae_le_trim, ae_le_trim_of_strongly_measurable hm hf hg⟩ + end integral_trim end measure_theory diff --git a/src/measure_theory/integral/set_integral.lean b/src/measure_theory/integral/set_integral.lean index 22888a802521f..b6d15435593f0 100644 --- a/src/measure_theory/integral/set_integral.lean +++ b/src/measure_theory/integral/set_integral.lean @@ -141,6 +141,13 @@ begin ... = ∫ x in s, f x ∂μ : by simp end +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 ∂μ := +by rw [← set.indicator_add_compl_eq_piecewise, + integral_add' (hf.indicator hs) (hg.indicator hs.compl), + integral_indicator hs, integral_indicator hs.compl] + lemma tendsto_set_integral_of_monotone {ι : Type*} [encodable ι] [semilattice_sup ι] {s : ι → set α} {f : α → E} (hsm : ∀ i, measurable_set (s i)) (h_mono : monotone s) (hfi : integrable_on f (⋃ n, s n) μ) : diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index 09372f3c467b0..1480575acbb02 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -3095,11 +3095,20 @@ lemma measure_trim_to_measurable_eq_zero {hm : m ≤ m0} (hs : μ.trim hm s = 0) μ (@to_measurable α m (μ.trim hm) s) = 0 := measure_eq_zero_of_trim_eq_zero hm (by rwa measure_to_measurable) +lemma ae_of_ae_trim (hm : m ≤ m0) {μ : measure α} {P : α → Prop} (h : ∀ᵐ x ∂(μ.trim hm), P x) : + ∀ᵐ x ∂μ, P x := +measure_eq_zero_of_trim_eq_zero hm h + lemma ae_eq_of_ae_eq_trim {E} {hm : m ≤ m0} {f₁ f₂ : α → E} (h12 : f₁ =ᶠ[@measure.ae α m (μ.trim hm)] f₂) : f₁ =ᵐ[μ] f₂ := measure_eq_zero_of_trim_eq_zero hm h12 +lemma ae_le_of_ae_le_trim {E} [has_le E] {hm : m ≤ m0} {f₁ f₂ : α → E} + (h12 : f₁ ≤ᶠ[@measure.ae α m (μ.trim hm)] f₂) : + f₁ ≤ᵐ[μ] f₂ := +measure_eq_zero_of_trim_eq_zero hm h12 + lemma trim_trim {m₁ m₂ : measurable_space α} {hm₁₂ : m₁ ≤ m₂} {hm₂ : m₂ ≤ m0} : (μ.trim hm₂).trim hm₁₂ = μ.trim (hm₁₂.trans hm₂) := begin diff --git a/src/probability/martingale.lean b/src/probability/martingale.lean index e9fe44bbf75ab..02b85d1b9e06b 100644 --- a/src/probability/martingale.lean +++ b/src/probability/martingale.lean @@ -225,6 +225,7 @@ begin simpa, end +/-- The converse of this lemma is `measure_theory.submartingale_of_set_integral_le`. -/ lemma set_integral_le {f : ι → α → ℝ} (hf : submartingale f ℱ μ) {i j : ι} (hij : i ≤ j) {s : set α} (hs : measurable_set[ℱ i] s) : ∫ x in s, f i x ∂μ ≤ ∫ x in s, f j x ∂μ := @@ -243,6 +244,31 @@ hf.sub_supermartingale hg.supermartingale end submartingale +section + +lemma submartingale_of_set_integral_le [is_finite_measure μ] + {f : ι → α → ℝ} (hadp : adapted ℱ f) (hint : ∀ i, integrable (f i) μ) + (hf : ∀ i j : ι, i ≤ j → ∀ s : set α, measurable_set[ℱ i] s → + ∫ x in s, f i x ∂μ ≤ ∫ x in s, f j x ∂μ) : + submartingale f ℱ μ := +begin + refine ⟨hadp, λ i j hij, _, hint⟩, + suffices : f i ≤ᵐ[μ.trim (ℱ.le i)] μ[f j| ℱ.le i], + { exact ae_le_of_ae_le_trim this }, + suffices : 0 ≤ᵐ[μ.trim (ℱ.le i)] μ[f j| ℱ.le i] - f i, + { filter_upwards [this] with x hx, + rwa ← sub_nonneg }, + refine ae_nonneg_of_forall_set_integral_nonneg_of_finite_measure + ((integrable_condexp.sub (hint i)).trim _ (strongly_measurable_condexp.sub $ hadp i)) + (λ s hs, _), + specialize hf i j hij s hs, + rwa [← set_integral_trim _ (strongly_measurable_condexp.sub $ hadp i) hs, + integral_sub' integrable_condexp.integrable_on (hint i).integrable_on, sub_nonneg, + set_integral_condexp _ (hint j) hs], +end + +end + namespace supermartingale lemma sub_submartingale [preorder E] [covariant_class E E (+) (≤)] @@ -353,6 +379,37 @@ end end submartingale +/-- The converse direction of the optional stopping theorem, i.e. an adapted integrable process `f` +is a submartingale if for all bounded stopping times `τ` and `π` such that `τ ≤ π`, the +stopped value of `f` at `τ` has expectation smaller than its stopped value at `π`. -/ +lemma submartingale_of_expected_stopped_value_mono [is_finite_measure μ] + {f : ℕ → α → ℝ} (hadp : adapted 𝒢 f) (hint : ∀ i, integrable (f i) μ) + (hf : ∀ τ π : α → ℕ, is_stopping_time 𝒢 τ → is_stopping_time 𝒢 π → τ ≤ π → (∃ N, ∀ x, π x ≤ N) → + μ[stopped_value f τ] ≤ μ[stopped_value f π]) : + submartingale f 𝒢 μ := +begin + refine submartingale_of_set_integral_le hadp hint (λ i j hij s hs, _), + classical, + specialize hf (s.piecewise (λ _, i) (λ _, j)) _ + (is_stopping_time_piecewise_const hij hs) + (is_stopping_time_const j) (λ x, (ite_le_sup _ _ _).trans (max_eq_right hij).le) + ⟨j, λ x, le_rfl⟩, + rwa [stopped_value_const, stopped_value_piecewise_const, + integral_piecewise (𝒢.le _ _ hs) (hint _).integrable_on (hint _).integrable_on, + ← integral_add_compl (𝒢.le _ _ hs) (hint j), add_le_add_iff_right] at hf, +end + +/-- **The optional stopping theorem** (fair game theorem): an adapted integrable process `f` +is a submartingale if and only if for all bounded stopping times `τ` and `π` such that `τ ≤ π`, the +stopped value of `f` at `τ` has expectation smaller than its stopped value at `π`. -/ +lemma submartingale_iff_expected_stopped_value_mono [is_finite_measure μ] + {f : ℕ → α → ℝ} (hadp : adapted 𝒢 f) (hint : ∀ i, integrable (f i) μ) : + submartingale f 𝒢 μ ↔ + ∀ τ π : α → ℕ, is_stopping_time 𝒢 τ → is_stopping_time 𝒢 π → τ ≤ π → (∃ N, ∀ x, π x ≤ N) → + μ[stopped_value f τ] ≤ μ[stopped_value f π] := +⟨λ hf _ _ hτ hπ hle ⟨N, hN⟩, hf.expected_stopped_value_mono hτ hπ hle hN, + submartingale_of_expected_stopped_value_mono hadp hint⟩ + end nat end measure_theory diff --git a/src/probability/stopping.lean b/src/probability/stopping.lean index 97497ec283991..5060720fd8d69 100644 --- a/src/probability/stopping.lean +++ b/src/probability/stopping.lean @@ -550,6 +550,38 @@ begin exact f.mono (sub_le_self j hi) _ (hτ (j - i)), end +lemma add_const_nat + {f : filtration ℕ m} {τ : α → ℕ} (hτ : is_stopping_time f τ) {i : ℕ} : + is_stopping_time f (λ x, τ x + i) := +begin + refine is_stopping_time_of_measurable_set_eq (λ j, _), + by_cases hij : i ≤ j, + { simp_rw [eq_comm, ← nat.sub_eq_iff_eq_add hij, eq_comm], + exact f.mono (j.sub_le i) _ (hτ.measurable_set_eq (j - i)) }, + { rw not_le at hij, + convert measurable_set.empty, + ext x, + simp only [set.mem_empty_eq, iff_false], + rintro (hx : τ x + i = j), + linarith }, +end + +-- generalize to certain encodable type? +lemma add + {f : filtration ℕ m} {τ π : α → ℕ} (hτ : is_stopping_time f τ) (hπ : is_stopping_time f π) : + is_stopping_time f (τ + π) := +begin + intro i, + rw (_ : {x | (τ + π) x ≤ i} = ⋃ k ≤ i, {x | π x = k} ∩ {x | τ x + k ≤ i}), + { exact measurable_set.Union (λ k, measurable_set.Union_Prop + (λ hk, (hπ.measurable_set_eq_le hk).inter (hτ.add_const_nat i))) }, + ext, + simp only [pi.add_apply, set.mem_set_of_eq, set.mem_Union, set.mem_inter_eq, exists_prop], + refine ⟨λ h, ⟨π x, by linarith, rfl, h⟩, _⟩, + rintro ⟨j, hj, rfl, h⟩, + assumption +end + section preorder variables [preorder ι] {f : filtration ι m} @@ -667,6 +699,9 @@ time `τ` is the map `x ↦ u (τ x) x`. -/ def stopped_value (u : ι → α → β) (τ : α → ι) : α → β := λ x, u (τ x) x +lemma stopped_value_const (u : ι → α → β) (i : ι) : stopped_value u (λ x, i) = u i := +rfl + variable [linear_order ι] /-- Given a map `u : ι → α → E`, the stopped process with respect to `τ` is `u i x` if @@ -897,4 +932,45 @@ end normed_group end nat +section piecewise_const + +variables [preorder ι] {𝒢 : filtration ι m} {τ η : α → ι} {i j : ι} {s : set α} + [decidable_pred (∈ s)] + +/-- Given stopping times `τ` and `η` which are bounded below, `set.piecewise s τ η` is also +a stopping time with respect to the same filtration. -/ +lemma is_stopping_time.piecewise_of_le (hτ_st : is_stopping_time 𝒢 τ) + (hη_st : is_stopping_time 𝒢 η) (hτ : ∀ x, i ≤ τ x) (hη : ∀ x, i ≤ η x) + (hs : measurable_set[𝒢 i] s) : + is_stopping_time 𝒢 (s.piecewise τ η) := +begin + intro n, + have : {x | s.piecewise τ η x ≤ n} + = (s ∩ {x | τ x ≤ n}) ∪ (sᶜ ∩ {x | η x ≤ n}), + { ext1 x, + simp only [set.piecewise, set.mem_inter_eq, set.mem_set_of_eq, and.congr_right_iff], + by_cases hx : x ∈ s; simp [hx], }, + rw this, + by_cases hin : i ≤ n, + { have hs_n : measurable_set[𝒢 n] s, from 𝒢.mono hin _ hs, + exact (hs_n.inter (hτ_st n)).union (hs_n.compl.inter (hη_st n)), }, + { have hτn : ∀ x, ¬ τ x ≤ n := λ x hτn, hin ((hτ x).trans hτn), + have hηn : ∀ x, ¬ η x ≤ n := λ x hηn, hin ((hη x).trans hηn), + simp [hτn, hηn], }, +end + +lemma is_stopping_time_piecewise_const (hij : i ≤ j) (hs : measurable_set[𝒢 i] s) : + is_stopping_time 𝒢 (s.piecewise (λ _, i) (λ _, j)) := +(is_stopping_time_const i).piecewise_of_le (is_stopping_time_const j) (λ x, le_rfl) (λ _, hij) hs + +lemma stopped_value_piecewise_const {ι' : Type*} {i j : ι'} {f : ι' → α → ℝ} : + stopped_value f (s.piecewise (λ _, i) (λ _, j)) = s.piecewise (f i) (f j) := +by { ext x, rw stopped_value, by_cases hx : x ∈ s; simp [hx] } + +lemma stopped_value_piecewise_const' {ι' : Type*} {i j : ι'} {f : ι' → α → ℝ} : + stopped_value f (s.piecewise (λ _, i) (λ _, j)) = s.indicator (f i) + sᶜ.indicator (f j) := +by { ext x, rw stopped_value, by_cases hx : x ∈ s; simp [hx] } + +end piecewise_const + end measure_theory From a6275694804455fe8995bd530e86b67ddab5cff1 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Mon, 2 May 2022 11:44:55 +0000 Subject: [PATCH 002/153] feat(category_theory/monoidal): adjunctions in rigid categories (#13707) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We construct the bijection on hom-sets `(Yᘁ ⊗ X ⟶ Z) ≃ (X ⟶ Y ⊗ Z)` given by "pulling the string on the left" down or up, using right duals in a right rigid category. As consequences, we show that a left rigid category is monoidal closed (it seems our lefts and rights have got mixed up!!), and that functors from a groupoid to a rigid category is again a rigid category. Co-authored-by: Scott Morrison --- src/algebra/category/FinVect.lean | 2 +- src/category_theory/monoidal/coherence.lean | 10 + .../monoidal/{rigid.lean => rigid/basic.lean} | 310 +++++++++++++++++- .../monoidal/rigid/functor_category.lean | 76 +++++ .../monoidal/rigid/of_equivalence.lean | 2 +- 5 files changed, 396 insertions(+), 4 deletions(-) rename src/category_theory/monoidal/{rigid.lean => rigid/basic.lean} (53%) create mode 100644 src/category_theory/monoidal/rigid/functor_category.lean diff --git a/src/algebra/category/FinVect.lean b/src/algebra/category/FinVect.lean index f573389dabe8b..7de3033e5d5f0 100644 --- a/src/algebra/category/FinVect.lean +++ b/src/algebra/category/FinVect.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Jakob von Raumer. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jakob von Raumer -/ -import category_theory.monoidal.rigid +import category_theory.monoidal.rigid.basic import linear_algebra.tensor_product_basis import linear_algebra.coevaluation import algebra.category.Module.monoidal diff --git a/src/category_theory/monoidal/coherence.lean b/src/category_theory/monoidal/coherence.lean index 4a6caca9a66b6..bc1e0775c9553 100644 --- a/src/category_theory/monoidal/coherence.lean +++ b/src/category_theory/monoidal/coherence.lean @@ -263,6 +263,11 @@ begin cases w, end +lemma insert_id_lhs {C : Type*} [category C] {X Y : C} (f g : X ⟶ Y) (w : f ≫ 𝟙 _ = g) : f = g := +by simpa using w +lemma insert_id_rhs {C : Type*} [category C] {X Y : C} (f g : X ⟶ Y) (w : f = g ≫ 𝟙 _) : f = g := +by simpa using w + end coherence open coherence @@ -300,6 +305,8 @@ do -- Then check that either `g₀` is identically `g₁`, reflexivity <|> (do -- or that both are compositions, + (do `(_ ≫ _ = _) ← target, skip) <|> `[apply tactic.coherence.insert_id_lhs], + (do `(_ = _ ≫ _) ← target, skip) <|> `[apply tactic.coherence.insert_id_rhs], `(_ ≫ _ = _ ≫ _) ← target | fail "`coherence` tactic failed, non-structural morphisms don't match", tactic.congr_core', @@ -323,6 +330,9 @@ example {U V W X Y : C} (f : U ⟶ V ⊗ (W ⊗ X)) (g : (V ⊗ W) ⊗ X ⟶ Y) f ⊗≫ g = f ≫ (α_ _ _ _).inv ≫ g := by coherence +example {U : C} (f : U ⟶ 𝟙_ C) : f ≫ (ρ_ (𝟙_ C)).inv ≫ (λ_ (𝟙_ C)).hom = f := +by coherence + example (W X Y Z : C) (f) : ((α_ W X Y).hom ⊗ 𝟙 Z) ≫ (α_ W (X ⊗ Y) Z).hom ≫ (𝟙 W ⊗ (α_ X Y Z).hom) ≫ f ≫ (α_ (W ⊗ X) Y Z).hom ≫ (α_ W X (Y ⊗ Z)).hom = diff --git a/src/category_theory/monoidal/rigid.lean b/src/category_theory/monoidal/rigid/basic.lean similarity index 53% rename from src/category_theory/monoidal/rigid.lean rename to src/category_theory/monoidal/rigid/basic.lean index ea296f09dde89..bb43e26be3124 100644 --- a/src/category_theory/monoidal/rigid.lean +++ b/src/category_theory/monoidal/rigid/basic.lean @@ -4,6 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jakob von Raumer -/ import category_theory.monoidal.coherence_lemmas +import category_theory.closed.monoidal +import tactic.apply_fun /-! # Rigid (autonomous) monoidal categories @@ -34,11 +36,17 @@ exact pairings and duals. * Show that `X ⊗ Y` and `Yᘁ ⊗ Xᘁ` form an exact pairing. * Show that the left adjoint mate of the right adjoint mate of a morphism is the morphism itself. * Simplify constructions in the case where a symmetry or braiding is present. -* Connect this definition to `monoidal_closed`: an object with a (left?) dual is - a closed object `X` such that the right adjoint of `X ⊗ -` is given by `Y ⊗ -` for some `Y`. * Show that `ᘁ` gives an equivalence of categories `C ≅ (Cᵒᵖ)ᴹᵒᵖ`. * Define pivotal categories (rigid categories equipped with a natural isomorphism `ᘁᘁ ≅ 𝟙 C`). +## Notes + +Although we construct the adjunction `tensor_left Y ⊣ tensor_left X` from `exact_pairing X Y`, +this is not a bijective correspondence. +I think the correct statement is that `tensor_left Y` and `tensor_left X` are +module endofunctors of `C` as a right `C` module category, +and `exact_pairing X Y` is in bijection with adjunctions compatible with this right `C` action. + ## References * @@ -227,6 +235,298 @@ begin right_unitor_naturality_assoc, ←unitors_equal, ←category.assoc, ←category.assoc], simp end +/-- +Given an exact pairing on `Y Y'`, +we get a bijection on hom-sets `(Y' ⊗ X ⟶ Z) ≃ (X ⟶ Y ⊗ Z)` +by "pulling the string on the left" up or down. + +This gives the adjunction `tensor_left_adjunction Y Y' : tensor_left Y' ⊣ tensor_left Y`. + +This adjunction is often referred to as "Frobenius reciprocity" in the +fusion categories / planar algebras / subfactors literature. +-/ +def tensor_left_hom_equiv (X Y Y' Z : C) [exact_pairing Y Y'] : + (Y' ⊗ X ⟶ Z) ≃ (X ⟶ Y ⊗ Z) := +{ to_fun := λ f, (λ_ _).inv ≫ (η_ _ _ ⊗ 𝟙 _) ≫ (α_ _ _ _).hom ≫ (𝟙 _ ⊗ f), + inv_fun := λ f, (𝟙 Y' ⊗ f) ≫ (α_ _ _ _).inv ≫ (ε_ _ _ ⊗ 𝟙 _) ≫ (λ_ _).hom, + left_inv := λ f, begin + dsimp, + simp only [id_tensor_comp], + slice_lhs 4 5 { rw associator_inv_naturality, }, + slice_lhs 5 6 { rw [tensor_id, id_tensor_comp_tensor_id, ←tensor_id_comp_id_tensor], }, + slice_lhs 2 5 { simp only [←tensor_id, associator_inv_conjugation], }, + have c : (α_ Y' (Y ⊗ Y') X).hom ≫ (𝟙 Y' ⊗ (α_ Y Y' X).hom) ≫ (α_ Y' Y (Y' ⊗ X)).inv ≫ + (α_ (Y' ⊗ Y) Y' X).inv = (α_ _ _ _).inv ⊗ 𝟙 _, pure_coherence, + slice_lhs 4 7 { rw c, }, + slice_lhs 3 5 { rw [←comp_tensor_id, ←comp_tensor_id, coevaluation_evaluation], }, + simp only [left_unitor_conjugation], + coherence, + end, + right_inv := λ f, begin + dsimp, + simp only [id_tensor_comp], + slice_lhs 3 4 { rw ←associator_naturality, }, + slice_lhs 2 3 { rw [tensor_id, tensor_id_comp_id_tensor, ←id_tensor_comp_tensor_id], }, + slice_lhs 3 6 { simp only [←tensor_id, associator_inv_conjugation], }, + have c : (α_ (Y ⊗ Y') Y Z).hom ≫ (α_ Y Y' (Y ⊗ Z)).hom ≫ (𝟙 Y ⊗ (α_ Y' Y Z).inv) ≫ + (α_ Y (Y' ⊗ Y) Z).inv = (α_ _ _ _).hom ⊗ 𝟙 Z, pure_coherence, + slice_lhs 5 8 { rw c, }, + slice_lhs 4 6 { rw [←comp_tensor_id, ←comp_tensor_id, evaluation_coevaluation], }, + simp only [left_unitor_conjugation], + coherence, + end, } + +/-- +Given an exact pairing on `Y Y'`, +we get a bijection on hom-sets `(X ⊗ Y ⟶ Z) ≃ (X ⟶ Z ⊗ Y')` +by "pulling the string on the right" up or down. +-/ +def tensor_right_hom_equiv (X Y Y' Z : C) [exact_pairing Y Y'] : + (X ⊗ Y ⟶ Z) ≃ (X ⟶ Z ⊗ Y') := +{ to_fun := λ f, (ρ_ _).inv ≫ (𝟙 _ ⊗ η_ _ _) ≫ (α_ _ _ _).inv ≫ (f ⊗ 𝟙 _), + inv_fun := λ f, (f ⊗ 𝟙 _) ≫ (α_ _ _ _).hom ≫ (𝟙 _ ⊗ ε_ _ _) ≫ (ρ_ _).hom, + left_inv := λ f, begin + dsimp, + simp only [comp_tensor_id], + slice_lhs 4 5 { rw associator_naturality, }, + slice_lhs 5 6 { rw [tensor_id, tensor_id_comp_id_tensor, ←id_tensor_comp_tensor_id], }, + slice_lhs 2 5 { simp only [←tensor_id, associator_conjugation], }, + have c : (α_ X (Y ⊗ Y') Y).inv ≫ ((α_ X Y Y').inv ⊗ 𝟙 Y) ≫ (α_ (X ⊗ Y) Y' Y).hom ≫ + (α_ X Y (Y' ⊗ Y)).hom = 𝟙 _ ⊗ (α_ _ _ _).hom, pure_coherence, + slice_lhs 4 7 { rw c, }, + slice_lhs 3 5 { rw [←id_tensor_comp, ←id_tensor_comp, evaluation_coevaluation], }, + simp only [right_unitor_conjugation], + coherence, + end, + right_inv := λ f, begin + dsimp, + simp only [comp_tensor_id], + slice_lhs 3 4 { rw ←associator_inv_naturality, }, + slice_lhs 2 3 { rw [tensor_id, id_tensor_comp_tensor_id, ←tensor_id_comp_id_tensor], }, + slice_lhs 3 6 { simp only [←tensor_id, associator_conjugation], }, + have c : (α_ Z Y' (Y ⊗ Y')).inv ≫ (α_ (Z ⊗ Y') Y Y').inv ≫ ((α_ Z Y' Y).hom ⊗ 𝟙 Y') ≫ + (α_ Z (Y' ⊗ Y) Y').hom = 𝟙 _ ⊗ (α_ _ _ _).inv, pure_coherence, + slice_lhs 5 8 { rw c, }, + slice_lhs 4 6 { rw [←id_tensor_comp, ←id_tensor_comp, coevaluation_evaluation], }, + simp only [right_unitor_conjugation], + coherence, + end, } + +lemma tensor_left_hom_equiv_naturality + {X Y Y' Z Z' : C} [exact_pairing Y Y'] (f : Y' ⊗ X ⟶ Z) (g : Z ⟶ Z') : + (tensor_left_hom_equiv X Y Y' Z') (f ≫ g) = + (tensor_left_hom_equiv X Y Y' Z) f ≫ (𝟙 Y ⊗ g) := +begin + dsimp [tensor_left_hom_equiv], + simp only [id_tensor_comp, category.assoc], +end + +lemma tensor_left_hom_equiv_symm_naturality {X X' Y Y' Z : C} [exact_pairing Y Y'] + (f : X ⟶ X') (g : X' ⟶ Y ⊗ Z) : + (tensor_left_hom_equiv X Y Y' Z).symm (f ≫ g) = + (𝟙 _ ⊗ f) ≫ (tensor_left_hom_equiv X' Y Y' Z).symm g := +begin + dsimp [tensor_left_hom_equiv], + simp only [id_tensor_comp, category.assoc], +end + +lemma tensor_right_hom_equiv_naturality {X Y Y' Z Z' : C} [exact_pairing Y Y'] + (f : X ⊗ Y ⟶ Z) (g : Z ⟶ Z') : + (tensor_right_hom_equiv X Y Y' Z') (f ≫ g) = + (tensor_right_hom_equiv X Y Y' Z) f ≫ (g ⊗ 𝟙 Y') := +begin + dsimp [tensor_right_hom_equiv], + simp only [comp_tensor_id, category.assoc], +end + +lemma tensor_right_hom_equiv_symm_naturality + {X X' Y Y' Z : C} [exact_pairing Y Y'] (f : X ⟶ X') (g : X' ⟶ Z ⊗ Y') : + ((tensor_right_hom_equiv X Y Y' Z).symm) (f ≫ g) = + (f ⊗ 𝟙 Y) ≫ ((tensor_right_hom_equiv X' Y Y' Z).symm) g := +begin + dsimp [tensor_right_hom_equiv], + simp only [comp_tensor_id, category.assoc], +end + +/-- +If `Y Y'` have an exact pairing, +then the functor `tensor_left Y'` is left adjoint to `tensor_left Y`. +-/ +def tensor_left_adjunction (Y Y' : C) [exact_pairing Y Y'] : tensor_left Y' ⊣ tensor_left Y := +adjunction.mk_of_hom_equiv +{ hom_equiv := λ X Z, tensor_left_hom_equiv X Y Y' Z, + hom_equiv_naturality_left_symm' := + λ X X' Z f g, tensor_left_hom_equiv_symm_naturality f g, + hom_equiv_naturality_right' := + λ X Z Z' f g, tensor_left_hom_equiv_naturality f g, } + +/-- +If `Y Y'` have an exact pairing, +then the functor `tensor_right Y` is left adjoint to `tensor_right Y'`. +-/ +def tensor_right_adjunction (Y Y' : C) [exact_pairing Y Y'] : tensor_right Y ⊣ tensor_right Y' := +adjunction.mk_of_hom_equiv +{ hom_equiv := λ X Z, tensor_right_hom_equiv X Y Y' Z, + hom_equiv_naturality_left_symm' := + λ X X' Z f g, tensor_right_hom_equiv_symm_naturality f g, + hom_equiv_naturality_right' := + λ X Z Z' f g, tensor_right_hom_equiv_naturality f g, } + +@[priority 100] +instance closed_of_has_left_dual (Y : C) [has_left_dual Y] : closed Y := +{ is_adj := ⟨_, tensor_left_adjunction (ᘁY) Y⟩, } + +/-- `tensor_left_hom_equiv` commutes with tensoring on the right -/ +lemma tensor_left_hom_equiv_tensor {X X' Y Y' Z Z' : C} [exact_pairing Y Y'] + (f : X ⟶ Y ⊗ Z) (g : X' ⟶ Z') : + (tensor_left_hom_equiv (X ⊗ X') Y Y' (Z ⊗ Z')).symm ((f ⊗ g) ≫ (α_ _ _ _).hom) = + (α_ _ _ _).inv ≫ ((tensor_left_hom_equiv X Y Y' Z).symm f ⊗ g) := +begin + dsimp [tensor_left_hom_equiv], + simp only [id_tensor_comp], + simp only [associator_inv_conjugation], + slice_lhs 2 2 { rw ←id_tensor_comp_tensor_id, }, + conv_rhs { rw [←id_tensor_comp_tensor_id, comp_tensor_id, comp_tensor_id], }, + simp, coherence, +end + +/-- `tensor_right_hom_equiv` commutes with tensoring on the left -/ +lemma tensor_right_hom_equiv_tensor {X X' Y Y' Z Z' : C} [exact_pairing Y Y'] + (f : X ⟶ Z ⊗ Y') (g : X' ⟶ Z') : + (tensor_right_hom_equiv (X' ⊗ X) Y Y' (Z' ⊗ Z)).symm ((g ⊗ f) ≫ (α_ _ _ _).inv) = + (α_ _ _ _).hom ≫ (g ⊗ (tensor_right_hom_equiv X Y Y' Z).symm f) := +begin + dsimp [tensor_right_hom_equiv], + simp only [comp_tensor_id], + simp only [associator_conjugation], + slice_lhs 2 2 { rw ←tensor_id_comp_id_tensor, }, + conv_rhs { rw [←tensor_id_comp_id_tensor, id_tensor_comp, id_tensor_comp], }, + simp only [←tensor_id, associator_conjugation], + simp, coherence, +end + +@[simp] +lemma tensor_left_hom_equiv_symm_coevaluation_comp_id_tensor + {Y Y' Z : C} [exact_pairing Y Y'] (f : Y' ⟶ Z) : + (tensor_left_hom_equiv _ _ _ _).symm (η_ _ _ ≫ (𝟙 Y ⊗ f)) = (ρ_ _).hom ≫ f := +begin + dsimp [tensor_left_hom_equiv], + rw id_tensor_comp, + slice_lhs 2 3 { rw associator_inv_naturality, }, + slice_lhs 3 4 { rw [tensor_id, id_tensor_comp_tensor_id, ←tensor_id_comp_id_tensor], }, + slice_lhs 1 3 { rw coevaluation_evaluation, }, + simp, +end + +@[simp] +lemma tensor_left_hom_equiv_symm_coevaluation_comp_tensor_id + {X Y : C} [has_right_dual X] [has_right_dual Y] (f : X ⟶ Y) : + (tensor_left_hom_equiv _ _ _ _).symm (η_ _ _ ≫ (f ⊗ 𝟙 Xᘁ)) = (ρ_ _).hom ≫ fᘁ := +begin + dsimp [tensor_left_hom_equiv, right_adjoint_mate], + simp, +end + +@[simp] +lemma tensor_right_hom_equiv_symm_coevaluation_comp_id_tensor + {X Y : C} [has_left_dual X] [has_left_dual Y] (f : X ⟶ Y) : + (tensor_right_hom_equiv _ (ᘁY) _ _).symm (η_ (ᘁX) X ≫ (𝟙 (ᘁX) ⊗ f)) = (λ_ _).hom ≫ (ᘁf) := +begin + dsimp [tensor_right_hom_equiv, left_adjoint_mate], + simp, +end + +@[simp] +lemma tensor_right_hom_equiv_symm_coevaluation_comp_tensor_id + {Y Y' Z : C} [exact_pairing Y Y'] (f : Y ⟶ Z) : + (tensor_right_hom_equiv _ Y _ _).symm (η_ Y Y' ≫ (f ⊗ 𝟙 Y')) = (λ_ _).hom ≫ f := +begin + dsimp [tensor_right_hom_equiv], + rw comp_tensor_id, + slice_lhs 2 3 { rw associator_naturality, }, + slice_lhs 3 4 { rw [tensor_id, tensor_id_comp_id_tensor, ←id_tensor_comp_tensor_id], }, + slice_lhs 1 3 { rw evaluation_coevaluation, }, + simp, +end + +@[simp] +lemma tensor_left_hom_equiv_id_tensor_comp_evaluation + {Y Z : C} [has_left_dual Z] (f : Y ⟶ (ᘁZ)) : + (tensor_left_hom_equiv _ _ _ _) ((𝟙 Z ⊗ f) ≫ ε_ _ _) = f ≫ (ρ_ _).inv := +begin + dsimp [tensor_left_hom_equiv], + rw id_tensor_comp, + slice_lhs 3 4 { rw ←associator_naturality, }, + slice_lhs 2 3 { rw [tensor_id, tensor_id_comp_id_tensor, ←id_tensor_comp_tensor_id], }, + slice_lhs 3 5 { rw evaluation_coevaluation, }, + simp, +end + +@[simp] +lemma tensor_left_hom_equiv_tensor_id_comp_evaluation + {X Y : C} [has_left_dual X] [has_left_dual Y] (f : X ⟶ Y) : + (tensor_left_hom_equiv _ _ _ _) ((f ⊗ 𝟙 _) ≫ ε_ _ _) = (ᘁf) ≫ (ρ_ _).inv := +begin + dsimp [tensor_left_hom_equiv, left_adjoint_mate], + simp, +end + +@[simp] +lemma tensor_right_hom_equiv_id_tensor_comp_evaluation + {X Y : C} [has_right_dual X] [has_right_dual Y] (f : X ⟶ Y) : + (tensor_right_hom_equiv _ _ _ _) ((𝟙 Yᘁ ⊗ f) ≫ ε_ _ _) = fᘁ ≫ (λ_ _).inv := +begin + dsimp [tensor_right_hom_equiv, right_adjoint_mate], + simp, +end + +@[simp] +lemma tensor_right_hom_equiv_tensor_id_comp_evaluation + {X Y : C} [has_right_dual X] (f : Y ⟶ Xᘁ) : + (tensor_right_hom_equiv _ _ _ _) ((f ⊗ 𝟙 X) ≫ ε_ X Xᘁ) = f ≫ (λ_ _).inv := +begin + dsimp [tensor_right_hom_equiv], + rw comp_tensor_id, + slice_lhs 3 4 { rw ←associator_inv_naturality, }, + slice_lhs 2 3 { rw [tensor_id, id_tensor_comp_tensor_id, ←tensor_id_comp_id_tensor], }, + slice_lhs 3 5 { rw coevaluation_evaluation, }, + simp, +end + +-- Next four lemmas passing `fᘁ` or `ᘁf` through (co)evaluations. + +lemma coevaluation_comp_right_adjoint_mate + {X Y : C} [has_right_dual X] [has_right_dual Y] (f : X ⟶ Y) : + η_ Y Yᘁ ≫ (𝟙 _ ⊗ fᘁ) = η_ _ _ ≫ (f ⊗ 𝟙 _) := +begin + apply_fun (tensor_left_hom_equiv _ Y Yᘁ _).symm, + simp, +end + +lemma left_adjoint_mate_comp_evaluation + {X Y : C} [has_left_dual X] [has_left_dual Y] (f : X ⟶ Y) : + (𝟙 X ⊗ (ᘁf)) ≫ ε_ _ _ = (f ⊗ 𝟙 _) ≫ ε_ _ _ := +begin + apply_fun (tensor_left_hom_equiv _ (ᘁX) X _), + simp, +end + +lemma coevaluation_comp_left_adjoint_mate + {X Y : C} [has_left_dual X] [has_left_dual Y] (f : X ⟶ Y) : + η_ (ᘁY) Y ≫ ((ᘁf) ⊗ 𝟙 Y) = η_ (ᘁX) X ≫ (𝟙 (ᘁX) ⊗ f) := +begin + apply_fun (tensor_right_hom_equiv _ (ᘁY) Y _).symm, + simp, +end + +lemma right_adjoint_mate_comp_evaluation + {X Y : C} [has_right_dual X] [has_right_dual Y] (f : X ⟶ Y) : + (fᘁ ⊗ 𝟙 X) ≫ ε_ X Xᘁ = (𝟙 Yᘁ ⊗ f) ≫ ε_ Y Yᘁ := +begin + apply_fun (tensor_right_hom_equiv _ X (Xᘁ) _), + simp, +end + /-- Transport an exact pairing across an isomorphism in the first argument. -/ def exact_pairing_congr_left {X X' Y : C} [exact_pairing X' Y] (i : X ≅ X') : exact_pairing X Y := { evaluation := (𝟙 Y ⊗ i.hom) ≫ ε_ _ _, @@ -318,6 +618,12 @@ class left_rigid_category (C : Type u) [category.{v} C] [monoidal_category.{v} C attribute [instance, priority 100] right_rigid_category.right_dual attribute [instance, priority 100] left_rigid_category.left_dual +@[priority 100] +instance monoidal_closed_of_left_rigid_category + (C : Type u) [category.{v} C] [monoidal_category.{v} C] [left_rigid_category C] : + monoidal_closed C := +{ closed' := λ X, by apply_instance, } + /-- A rigid monoidal category is a monoidal category which is left rigid and right rigid. -/ class rigid_category (C : Type u) [category.{v} C] [monoidal_category.{v} C] extends right_rigid_category C, left_rigid_category C diff --git a/src/category_theory/monoidal/rigid/functor_category.lean b/src/category_theory/monoidal/rigid/functor_category.lean new file mode 100644 index 0000000000000..ef3315f137b3d --- /dev/null +++ b/src/category_theory/monoidal/rigid/functor_category.lean @@ -0,0 +1,76 @@ +/- +Copyright (c) 2022 Scott Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Scott Morrison +-/ +import category_theory.monoidal.rigid.basic +import category_theory.monoidal.functor_category + +/-! +# Functors from a groupoid into a right/left rigid category form a right/left rigid category. + +(Using the pointwise monoidal structure on the functor category.) +-/ + +noncomputable theory + +open category_theory +open category_theory.monoidal_category + +namespace category_theory.monoidal + +variables {C D : Type*} [groupoid C] [category D] [monoidal_category D] + +instance functor_has_right_dual [right_rigid_category D] (F : C ⥤ D) : has_right_dual F := +{ right_dual := + { obj := λ X, (F.obj X)ᘁ, + map := λ X Y f, (F.map (inv f))ᘁ, + map_comp' := λ X Y Z f g, by { simp [comp_right_adjoint_mate], }, }, + exact := + { evaluation := + { app := λ X, ε_ _ _, + naturality' := λ X Y f, begin + dsimp, + rw [category.comp_id, functor.map_inv, + ←id_tensor_comp_tensor_id, category.assoc, right_adjoint_mate_comp_evaluation, + ←category.assoc, ←id_tensor_comp, is_iso.hom_inv_id, tensor_id, category.id_comp], + end }, + coevaluation := + { app := λ X, η_ _ _, + naturality' := λ X Y f, begin + dsimp, + rw [functor.map_inv, category.id_comp, + ←id_tensor_comp_tensor_id, ←category.assoc, coevaluation_comp_right_adjoint_mate, + category.assoc, ←comp_tensor_id, is_iso.inv_hom_id, tensor_id, category.comp_id], + end, }, }, } + +instance right_rigid_functor_category [right_rigid_category D] : right_rigid_category (C ⥤ D) := {} + +instance functor_has_left_dual [left_rigid_category D] (F : C ⥤ D) : has_left_dual F := +{ left_dual := + { obj := λ X, ᘁ(F.obj X), + map := λ X Y f, ᘁ(F.map (inv f)), + map_comp' := λ X Y Z f g, by { simp [comp_left_adjoint_mate], }, }, + exact := + { evaluation := + { app := λ X, ε_ _ _, + naturality' := λ X Y f, begin + dsimp, + rw [category.comp_id, functor.map_inv, + ←tensor_id_comp_id_tensor, category.assoc, left_adjoint_mate_comp_evaluation, + ←category.assoc, ←comp_tensor_id, is_iso.hom_inv_id, tensor_id, category.id_comp], + end }, + coevaluation := + { app := λ X, η_ _ _, + naturality' := λ X Y f, begin + dsimp, + rw [functor.map_inv, category.id_comp, + ←tensor_id_comp_id_tensor, ←category.assoc, coevaluation_comp_left_adjoint_mate, + category.assoc, ←id_tensor_comp, is_iso.inv_hom_id, tensor_id, category.comp_id], + end, }, }, } + +instance left_rigid_functor_category [left_rigid_category D] : left_rigid_category (C ⥤ D) := {} + +instance rigid_functor_category [rigid_category D] : rigid_category (C ⥤ D) := {} + +end category_theory.monoidal diff --git a/src/category_theory/monoidal/rigid/of_equivalence.lean b/src/category_theory/monoidal/rigid/of_equivalence.lean index 4a67599bd27d3..567fa184962d7 100644 --- a/src/category_theory/monoidal/rigid/of_equivalence.lean +++ b/src/category_theory/monoidal/rigid/of_equivalence.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Scott Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ -import category_theory.monoidal.rigid +import category_theory.monoidal.rigid.basic /-! # Transport rigid structures over a monoidal equivalence. From 320df450e9abeb5fc6417971e75acb6ae8bc3794 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 2 May 2022 11:44:56 +0000 Subject: [PATCH 003/153] refactor(linear_algebra/trace): unbundle `matrix.trace` (#13712) These extra type arguments are annoying to work with in many cases, especially when Lean doesn't have any information to infer the mostly-irrelevant `R` argument from. This came up while trying to work with `continuous.matrix_trace`, which is annoying to use for that reason. The old bundled version is still available as `matrix.trace_linear_map`. The cost of this change is that we have to copy across the usual set of obvious lemmas about additive maps; but we already do this for `diagonal`, `transpose` etc anyway. --- src/algebra/lie/classical.lean | 10 +- .../simple_graph/adj_matrix.lean | 6 +- src/data/matrix/basic.lean | 12 ++ src/data/matrix/basis.lean | 4 +- src/data/matrix/hadamard.lean | 4 +- src/linear_algebra/matrix/charpoly/coeff.lean | 15 +- src/linear_algebra/matrix/trace.lean | 128 ++++++++++++------ src/linear_algebra/trace.lean | 22 +-- src/ring_theory/trace.lean | 8 +- src/topology/instances/matrix.lean | 6 +- 10 files changed, 138 insertions(+), 77 deletions(-) diff --git a/src/algebra/lie/classical.lean b/src/algebra/lie/classical.lean index 90cf7ae9b6cab..0bc4ada4aecb4 100644 --- a/src/algebra/lie/classical.lean +++ b/src/algebra/lie/classical.lean @@ -74,9 +74,9 @@ variables [decidable_eq n] [decidable_eq p] [decidable_eq q] [decidable_eq l] variables [comm_ring R] @[simp] lemma matrix_trace_commutator_zero [fintype n] (X Y : matrix n n R) : - matrix.trace n R R ⁅X, Y⁆ = 0 := -calc _ = matrix.trace n R R (X ⬝ Y) - matrix.trace n R R (Y ⬝ X) : linear_map.map_sub _ _ _ - ... = matrix.trace n R R (X ⬝ Y) - matrix.trace n R R (X ⬝ Y) : + matrix.trace ⁅X, Y⁆ = 0 := +calc _ = matrix.trace (X ⬝ Y) - matrix.trace (Y ⬝ X) : trace_sub _ _ + ... = matrix.trace (X ⬝ Y) - matrix.trace (X ⬝ Y) : congr_arg (λ x, _ - x) (matrix.trace_mul_comm Y X) ... = 0 : sub_self _ @@ -85,7 +85,7 @@ namespace special_linear /-- The special linear Lie algebra: square matrices of trace zero. -/ def sl [fintype n] : lie_subalgebra R (matrix n n R) := { lie_mem' := λ X Y _ _, linear_map.mem_ker.2 $ matrix_trace_commutator_zero _ _ _ _, - ..linear_map.ker (matrix.trace n R R) } + ..linear_map.ker (matrix.trace_linear_map n R R) } lemma sl_bracket [fintype n] (A B : sl n R) : ⁅A, B⁆.val = A.val ⬝ B.val - B.val ⬝ A.val := rfl @@ -97,7 +97,7 @@ variables {n} [fintype n] (i j : n) basis of sl n R. -/ def Eb (h : j ≠ i) : sl n R := ⟨matrix.std_basis_matrix i j (1 : R), - show matrix.std_basis_matrix i j (1 : R) ∈ linear_map.ker (matrix.trace n R R), + show matrix.std_basis_matrix i j (1 : R) ∈ linear_map.ker (matrix.trace_linear_map n R R), from matrix.std_basis_matrix.trace_zero i j (1 : R) h⟩ @[simp] lemma Eb_val (h : j ≠ i) : (Eb R i j h).val = matrix.std_basis_matrix i j 1 := rfl diff --git a/src/combinatorics/simple_graph/adj_matrix.lean b/src/combinatorics/simple_graph/adj_matrix.lean index 1269203b6d7b7..7ffea9676d53c 100644 --- a/src/combinatorics/simple_graph/adj_matrix.lean +++ b/src/combinatorics/simple_graph/adj_matrix.lean @@ -223,9 +223,9 @@ by simp [mul_apply, neighbor_finset_eq_filter, sum_filter, adj_comm] variable (α) -theorem trace_adj_matrix [non_assoc_semiring α] [semiring β] [module β α]: - matrix.trace _ β _ (G.adj_matrix α) = 0 := -by simp +@[simp] theorem trace_adj_matrix [add_comm_monoid α] [has_one α] : + matrix.trace (G.adj_matrix α) = 0 := +by simp [matrix.trace] variable {α} diff --git a/src/data/matrix/basic.lean b/src/data/matrix/basic.lean index 31f6f6b2cd021..443f84a8efc68 100644 --- a/src/data/matrix/basic.lean +++ b/src/data/matrix/basic.lean @@ -368,6 +368,18 @@ lemma diag_map {f : α → β} {A : matrix n n α} : diag (A.map f) = f ∘ diag @[simp] lemma diag_conj_transpose [add_monoid α] [star_add_monoid α] (A : matrix n n α) : diag Aᴴ = star (diag A) := rfl +@[simp] lemma diag_list_sum [add_monoid α] (l : list (matrix n n α)) : + diag l.sum = (l.map diag).sum := +map_list_sum (diag_add_monoid_hom n α) l + +@[simp] lemma diag_multiset_sum [add_comm_monoid α] (s : multiset (matrix n n α)) : + diag s.sum = (s.map diag).sum := +map_multiset_sum (diag_add_monoid_hom n α) s + +@[simp] lemma diag_sum {ι} [add_comm_monoid α] (s : finset ι) (f : ι → matrix n n α) : + diag (∑ i in s, f i) = ∑ i in s, diag (f i) := +map_sum (diag_add_monoid_hom n α) f s + end diag section dot_product diff --git a/src/data/matrix/basis.lean b/src/data/matrix/basis.lean index f60a9942cd9b7..cb40e55b8f0b4 100644 --- a/src/data/matrix/basis.lean +++ b/src/data/matrix/basis.lean @@ -129,7 +129,9 @@ by { ext j, by_cases hij : i = j; try {rw hij}; simp [hij] } variable [fintype n] -lemma trace_zero (h : j ≠ i) : trace n α α (std_basis_matrix i j c) = 0 := by simp [h] +@[simp] lemma trace_zero (h : j ≠ i) : trace (std_basis_matrix i j c) = 0 := by simp [trace, h] + +@[simp] lemma trace_eq : trace (std_basis_matrix i i c) = c := by simp [trace] @[simp] lemma mul_left_apply_same (b : n) (M : matrix n n α) : (std_basis_matrix i j c ⬝ M) i b = c * M j b := diff --git a/src/data/matrix/hadamard.lean b/src/data/matrix/hadamard.lean index 6361991da570f..2985d309e17d2 100644 --- a/src/data/matrix/hadamard.lean +++ b/src/data/matrix/hadamard.lean @@ -115,11 +115,11 @@ section trace variables [fintype m] [fintype n] variables (R) [semiring α] [semiring R] [module R α] -lemma sum_hadamard_eq : ∑ (i : m) (j : n), (A ⊙ B) i j = trace m R α (A ⬝ Bᵀ) := +lemma sum_hadamard_eq : ∑ (i : m) (j : n), (A ⊙ B) i j = trace (A ⬝ Bᵀ) := rfl lemma dot_product_vec_mul_hadamard [decidable_eq m] [decidable_eq n] (v : m → α) (w : n → α) : - dot_product (vec_mul v (A ⊙ B)) w = trace m R α (diagonal v ⬝ A ⬝ (B ⬝ diagonal w)ᵀ) := + dot_product (vec_mul v (A ⊙ B)) w = trace (diagonal v ⬝ A ⬝ (B ⬝ diagonal w)ᵀ) := begin rw [←sum_hadamard_eq, finset.sum_comm], simp [dot_product, vec_mul, finset.sum_mul, mul_assoc], diff --git a/src/linear_algebra/matrix/charpoly/coeff.lean b/src/linear_algebra/matrix/charpoly/coeff.lean index 7db25d01c7c70..afe8e5c24be35 100644 --- a/src/linear_algebra/matrix/charpoly/coeff.lean +++ b/src/linear_algebra/matrix/charpoly/coeff.lean @@ -118,11 +118,12 @@ begin end theorem trace_eq_neg_charpoly_coeff [nonempty n] (M : matrix n n R) : - (trace n R R) M = -M.charpoly.coeff (fintype.card n - 1) := + trace M = -M.charpoly.coeff (fintype.card n - 1) := begin rw charpoly_coeff_eq_prod_coeff_of_le, swap, refl, - rw [fintype.card, prod_X_sub_C_coeff_card_pred univ (λ i : n, M i i)], simp, - rw [← fintype.card, fintype.card_pos_iff], apply_instance, + rw [fintype.card, prod_X_sub_C_coeff_card_pred univ (λ i : n, M i i) fintype.card_pos, neg_neg, + trace], + refl end -- I feel like this should use polynomial.alg_hom_eval₂_algebra_map @@ -209,16 +210,16 @@ end by { have h := finite_field.matrix.charpoly_pow_card M, rwa zmod.card at h, } lemma finite_field.trace_pow_card {K : Type*} [field K] [fintype K] - (M : matrix n n K) : trace n K K (M ^ (fintype.card K)) = (trace n K K M) ^ (fintype.card K) := + (M : matrix n n K) : trace (M ^ (fintype.card K)) = trace M ^ (fintype.card K) := begin casesI is_empty_or_nonempty n, - { simp [zero_pow fintype.card_pos], }, + { simp [zero_pow fintype.card_pos, matrix.trace], }, rw [matrix.trace_eq_neg_charpoly_coeff, matrix.trace_eq_neg_charpoly_coeff, finite_field.matrix.charpoly_pow_card, finite_field.pow_card] end -lemma zmod.trace_pow_card {p:ℕ} [fact p.prime] (M : matrix n n (zmod p)) : - trace n (zmod p) (zmod p) (M ^ p) = (trace n (zmod p) (zmod p) M)^p := +lemma zmod.trace_pow_card {p : ℕ} [fact p.prime] (M : matrix n n (zmod p)) : + trace (M ^ p) = (trace M)^p := by { have h := finite_field.trace_pow_card M, rwa zmod.card at h, } namespace matrix diff --git a/src/linear_algebra/matrix/trace.lean b/src/linear_algebra/matrix/trace.lean index 5d8f912b3300f..fa0e18ef89d24 100644 --- a/src/linear_algebra/matrix/trace.lean +++ b/src/linear_algebra/matrix/trace.lean @@ -8,8 +8,8 @@ import data.matrix.basic /-! # Trace of a matrix -This file defines the trace of a matrix, the linear map -sending a matrix to the sum of its diagonal entries. +This file defines the trace of a matrix, the map sending a matrix to the sum of its diagonal +entries. See also `linear_algebra.trace` for the trace of an endomorphism. @@ -19,60 +19,106 @@ matrix, trace, diagonal -/ -open_locale big_operators -open_locale matrix +open_locale big_operators matrix namespace matrix -section trace +variables {ι m n p : Type*} {α R S : Type*} +variables [fintype m] [fintype n] [fintype p] -universes u v w +section add_comm_monoid +variables [add_comm_monoid R] -variables {m : Type*} (n : Type*) {p : Type*} -variables (R : Type*) (M : Type*) [semiring R] [add_comm_monoid M] [module R M] +/-- The trace of a square matrix. For more bundled versions, see: +* `matrix.trace_add_monoid_hom` +* `matrix.trace_linear_map` +-/ +def trace (A : matrix n n R) : R := ∑ i, diag A i -variables (n) (R) (M) +variables (n R) +@[simp] lemma trace_zero : trace (0 : matrix n n R) = 0 := +(finset.sum_const (0 : R)).trans $ smul_zero _ +variables {n R} -/-- -The trace of a square matrix. --/ -def trace [fintype n] : (matrix n n M) →ₗ[R] M := -{ to_fun := λ A, ∑ i, diag A i, - map_add' := by { intros, apply finset.sum_add_distrib, }, - map_smul' := by { intros, simp [finset.smul_sum], } } +@[simp] lemma trace_add (A B : matrix n n R) : trace (A + B) = trace A + trace B := +finset.sum_add_distrib + +@[simp] lemma trace_smul [monoid α] [distrib_mul_action α R] (r : α) (A : matrix n n R) : + trace (r • A) = r • trace A := +finset.smul_sum.symm + +@[simp] lemma trace_transpose (A : matrix n n R) : trace Aᵀ = trace A := rfl + +variables (n α R) +/-- `matrix.trace` as an `add_monoid_hom` -/ +@[simps] +def trace_add_monoid_hom : matrix n n R →+ R := +{ to_fun := trace, map_zero' := trace_zero n R, map_add' := trace_add } + +/-- `matrix.trace` as a `linear_map` -/ +@[simps] +def trace_linear_map [semiring α] [module α R] : matrix n n R →ₗ[α] R := +{ to_fun := trace, map_add' := trace_add, map_smul' := trace_smul } +variables {n α R} + +@[simp] lemma trace_list_sum (l : list (matrix n n R)) : trace l.sum = (l.map trace).sum := +map_list_sum (trace_add_monoid_hom n R) l + +@[simp] lemma trace_multiset_sum (s : multiset (matrix n n R)) : trace s.sum = (s.map trace).sum := +map_multiset_sum (trace_add_monoid_hom n R) s -variables {n} {R} {M} [fintype n] [fintype m] [fintype p] +@[simp] lemma trace_sum (s : finset ι) (f : ι → matrix n n R) : + trace (∑ i in s, f i) = ∑ i in s, trace (f i) := +map_sum (trace_add_monoid_hom n R) f s -@[simp] lemma trace_diag (A : matrix n n M) : trace n R M A = ∑ i, diag A i := rfl +end add_comm_monoid -lemma trace_apply (A : matrix n n M) : trace n R M A = ∑ i, A i i := rfl +section add_comm_group +variables [add_comm_group R] -@[simp] lemma trace_one [decidable_eq n] : - trace n R R 1 = fintype.card n := -have h : trace n R R 1 = ∑ i, diag 1 i := rfl, -by simp_rw [h, diag_one, pi.one_def, finset.sum_const, nsmul_one]; refl +@[simp] lemma trace_sub (A B : matrix n n R) : trace (A - B) = trace A - trace B := +finset.sum_sub_distrib -@[simp] lemma trace_transpose (A : matrix n n M) : trace n R M Aᵀ = trace n R M A := rfl +@[simp] lemma trace_neg (A : matrix n n R) : trace (-A) = -trace A := +finset.sum_neg_distrib -@[simp] lemma trace_transpose_mul (A : matrix m n R) (B : matrix n m R) : - trace n R R (Aᵀ ⬝ Bᵀ) = trace m R R (A ⬝ B) := finset.sum_comm +end add_comm_group -lemma trace_mul_comm {S : Type v} [comm_semiring S] (A : matrix m n S) (B : matrix n m S) : - trace m S S (A ⬝ B) = trace n S S (B ⬝ A) := +section one +variables [decidable_eq n] [add_comm_monoid R] [has_one R] + +@[simp] lemma trace_one : trace (1 : matrix n n R) = fintype.card n := +by simp_rw [trace, diag_one, pi.one_def, finset.sum_const, nsmul_one, finset.card_univ] + +end one + +section mul + +@[simp] lemma trace_transpose_mul [add_comm_monoid R] [has_mul R] + (A : matrix m n R) (B : matrix n m R) : trace (Aᵀ ⬝ Bᵀ) = trace (A ⬝ B) := finset.sum_comm + +lemma trace_mul_comm [add_comm_monoid R] [comm_semigroup R] (A : matrix m n R) (B : matrix n m R) : + trace (A ⬝ B) = trace (B ⬝ A) := by rw [←trace_transpose, ←trace_transpose_mul, transpose_mul] -lemma trace_mul_cycle {S : Type v} [comm_semiring S] - (A : matrix m n S) (B : matrix n p S) (C : matrix p m S) : - trace _ S S (A ⬝ B ⬝ C) = trace p S S (C ⬝ A ⬝ B) := +lemma trace_mul_cycle [non_unital_comm_semiring R] + (A : matrix m n R) (B : matrix n p R) (C : matrix p m R) : + trace (A ⬝ B ⬝ C) = trace (C ⬝ A ⬝ B) := by rw [trace_mul_comm, matrix.mul_assoc] -lemma trace_mul_cycle' {S : Type v} [comm_semiring S] - (A : matrix m n S) (B : matrix n p S) (C : matrix p m S) : - trace _ S S (A ⬝ (B ⬝ C)) = trace p S S (C ⬝ (A ⬝ B)) := +lemma trace_mul_cycle' [non_unital_comm_semiring R] + (A : matrix m n R) (B : matrix n p R) (C : matrix p m R) : + trace (A ⬝ (B ⬝ C)) = trace (C ⬝ (A ⬝ B)) := by rw [←matrix.mul_assoc, trace_mul_comm] -@[simp] lemma trace_col_mul_row (a b : n → R) : trace n R R (col a ⬝ row b) = dot_product a b := -by simp [dot_product] +@[simp] lemma trace_col_mul_row [non_unital_non_assoc_semiring R] (a b : n → R) : + trace (col a ⬝ row b) = dot_product a b := +by simp [dot_product, trace] + +end mul + +section fin +variables [add_comm_monoid R] /-! ### Special cases for `fin n` @@ -80,18 +126,18 @@ While `simp [fin.sum_univ_succ]` can prove these, we include them for convenienc with `matrix.det_fin_two` etc. -/ -@[simp] lemma trace_fin_zero (A : matrix (fin 0) (fin 0) R) : trace _ R R A = 0 := +@[simp] lemma trace_fin_zero (A : matrix (fin 0) (fin 0) R) : trace A = 0 := rfl -lemma trace_fin_one (A : matrix (fin 1) (fin 1) R) : trace _ R R A = A 0 0 := +lemma trace_fin_one (A : matrix (fin 1) (fin 1) R) : trace A = A 0 0 := add_zero _ -lemma trace_fin_two (A : matrix (fin 2) (fin 2) R) : trace _ R R A = A 0 0 + A 1 1 := +lemma trace_fin_two (A : matrix (fin 2) (fin 2) R) : trace A = A 0 0 + A 1 1 := congr_arg ((+) _) (add_zero (A 1 1)) -lemma trace_fin_three (A : matrix (fin 3) (fin 3) R) : trace _ R R A = A 0 0 + A 1 1 + A 2 2 := +lemma trace_fin_three (A : matrix (fin 3) (fin 3) R) : trace A = A 0 0 + A 1 1 + A 2 2 := by { rw [← add_zero (A 2 2), add_assoc], refl } -end trace +end fin end matrix diff --git a/src/linear_algebra/trace.lean b/src/linear_algebra/trace.lean index afbb55aa628bb..d6f9c100b8042 100644 --- a/src/linear_algebra/trace.lean +++ b/src/linear_algebra/trace.lean @@ -41,29 +41,29 @@ variables (b : basis ι R M) (c : basis κ R M) /-- The trace of an endomorphism given a basis. -/ def trace_aux : (M →ₗ[R] M) →ₗ[R] R := -(matrix.trace ι R R) ∘ₗ ↑(linear_map.to_matrix b b) +(matrix.trace_linear_map ι R R) ∘ₗ ↑(linear_map.to_matrix b b) -- Can't be `simp` because it would cause a loop. lemma trace_aux_def (b : basis ι R M) (f : M →ₗ[R] M) : - trace_aux R b f = matrix.trace ι R R (linear_map.to_matrix b b f) := + trace_aux R b f = matrix.trace (linear_map.to_matrix b b f) := rfl theorem trace_aux_eq : trace_aux R b = trace_aux R c := linear_map.ext $ λ f, -calc matrix.trace ι R R (linear_map.to_matrix b b f) - = matrix.trace ι R R (linear_map.to_matrix b b ((linear_map.id.comp f).comp linear_map.id)) : +calc matrix.trace (linear_map.to_matrix b b f) + = matrix.trace (linear_map.to_matrix b b ((linear_map.id.comp f).comp linear_map.id)) : by rw [linear_map.id_comp, linear_map.comp_id] -... = matrix.trace ι R R (linear_map.to_matrix c b linear_map.id ⬝ +... = matrix.trace (linear_map.to_matrix c b linear_map.id ⬝ linear_map.to_matrix c c f ⬝ linear_map.to_matrix b c linear_map.id) : by rw [linear_map.to_matrix_comp _ c, linear_map.to_matrix_comp _ c] -... = matrix.trace κ R R (linear_map.to_matrix c c f ⬝ +... = matrix.trace (linear_map.to_matrix c c f ⬝ linear_map.to_matrix b c linear_map.id ⬝ linear_map.to_matrix c b linear_map.id) : by rw [matrix.mul_assoc, matrix.trace_mul_comm] -... = matrix.trace κ R R (linear_map.to_matrix c c ((f.comp linear_map.id).comp linear_map.id)) : +... = matrix.trace (linear_map.to_matrix c c ((f.comp linear_map.id).comp linear_map.id)) : by rw [linear_map.to_matrix_comp _ b, linear_map.to_matrix_comp _ c] -... = matrix.trace κ R R (linear_map.to_matrix c c f) : +... = matrix.trace (linear_map.to_matrix c c f) : by rw [linear_map.comp_id, linear_map.comp_id] open_locale classical @@ -81,13 +81,13 @@ variables (R) {M} /-- Auxiliary lemma for `trace_eq_matrix_trace`. -/ theorem trace_eq_matrix_trace_of_finset {s : finset M} (b : basis s R M) (f : M →ₗ[R] M) : - trace R M f = matrix.trace s R R (linear_map.to_matrix b b f) := + trace R M f = matrix.trace (linear_map.to_matrix b b f) := have ∃ (s : finset M), nonempty (basis s R M), from ⟨s, ⟨b⟩⟩, by { rw [trace, dif_pos this, ← trace_aux_def], congr' 1, apply trace_aux_eq } theorem trace_eq_matrix_trace (f : M →ₗ[R] M) : - trace R M f = matrix.trace ι R R (linear_map.to_matrix b b f) := + trace R M f = matrix.trace (linear_map.to_matrix b b f) := by rw [trace_eq_matrix_trace_of_finset R b.reindex_finset_range, ← trace_aux_def, ← trace_aux_def, trace_aux_eq R b] @@ -127,7 +127,7 @@ begin simp only [function.comp_app, basis.tensor_product_apply, basis.coe_dual_basis, coe_comp], rw [trace_eq_matrix_trace R b, to_matrix_dual_tensor_hom], by_cases hij : i = j, - { rw [hij], simp}, + { rw [hij], simp }, rw matrix.std_basis_matrix.trace_zero j i (1:R) hij, simp [finsupp.single_eq_pi_single, hij], end diff --git a/src/ring_theory/trace.lean b/src/ring_theory/trace.lean index 657b938183d9e..69651a2c3beb4 100644 --- a/src/ring_theory/trace.lean +++ b/src/ring_theory/trace.lean @@ -98,7 +98,7 @@ variables {R} -- Can't be a `simp` lemma because it depends on a choice of basis lemma trace_eq_matrix_trace [decidable_eq ι] (b : basis ι R S) (s : S) : - trace R S s = matrix.trace _ R _ (algebra.left_mul_matrix b s) := + trace R S s = matrix.trace (algebra.left_mul_matrix b s) := by rw [trace_apply, linear_map.trace_eq_matrix_trace _ b, to_matrix_lmul_eq] /-- If `x` is in the base field `K`, then the trace is `[L : K] * x`. -/ @@ -106,7 +106,7 @@ lemma trace_algebra_map_of_basis (x : R) : trace R S (algebra_map R S x) = fintype.card ι • x := begin haveI := classical.dec_eq ι, - rw [trace_apply, linear_map.trace_eq_matrix_trace R b, trace_diag], + rw [trace_apply, linear_map.trace_eq_matrix_trace R b, matrix.trace], convert finset.sum_const _, ext i, simp, @@ -133,10 +133,10 @@ begin haveI := classical.dec_eq ι, haveI := classical.dec_eq κ, rw [trace_eq_matrix_trace (b.smul c), trace_eq_matrix_trace b, trace_eq_matrix_trace c, - matrix.trace_apply, matrix.trace_apply, matrix.trace_apply, + matrix.trace, matrix.trace, matrix.trace, ← finset.univ_product_univ, finset.sum_product], refine finset.sum_congr rfl (λ i _, _), - simp only [alg_hom.map_sum, smul_left_mul_matrix, finset.sum_apply, + simp only [alg_hom.map_sum, smul_left_mul_matrix, finset.sum_apply, matrix.diag, -- The unifier is not smart enough to apply this one by itself: finset.sum_apply i _ (λ y, left_mul_matrix b (left_mul_matrix c x y y))] end diff --git a/src/topology/instances/matrix.lean b/src/topology/instances/matrix.lean index e6d49b1850428..7a039af9faac6 100644 --- a/src/topology/instances/matrix.lean +++ b/src/topology/instances/matrix.lean @@ -159,9 +159,9 @@ lemma continuous_matrix_diag : continuous (matrix.diag : matrix n n R → n → show continuous (λ x : matrix n n R, matrix.diag x), from continuous_id.matrix_diag @[continuity] -lemma continuous.matrix_trace [fintype n] [semiring S] [add_comm_monoid R] [has_continuous_add R] - [module S R] {A : X → matrix n n R} (hA : continuous A) : - continuous (λ x, trace n S R (A x)) := +lemma continuous.matrix_trace [fintype n] [add_comm_monoid R] [has_continuous_add R] + {A : X → matrix n n R} (hA : continuous A) : + continuous (λ x, trace (A x)) := continuous_finset_sum _ $ λ i hi, hA.matrix_elem _ _ @[continuity] From 61d5d308c6b2268799d2e4377a6232b34aced678 Mon Sep 17 00:00:00 2001 From: tb65536 Date: Mon, 2 May 2022 11:44:57 +0000 Subject: [PATCH 004/153] feat(group_theory/group_action/basic): A multiplicative action induces an additive action of the additive group (#13780) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit `mul_action M α` induces `add_action (additive M) α`. --- src/group_theory/group_action/defs.lean | 19 +++++++++++++++++++ 1 file changed, 19 insertions(+) diff --git a/src/group_theory/group_action/defs.lean b/src/group_theory/group_action/defs.lean index e3845ecac017e..232be4278097b 100644 --- a/src/group_theory/group_action/defs.lean +++ b/src/group_theory/group_action/defs.lean @@ -85,6 +85,25 @@ class mul_action (α : Type*) (β : Type*) [monoid α] extends has_scalar α β (one_smul : ∀ b : β, (1 : α) • b = b) (mul_smul : ∀ (x y : α) (b : β), (x * y) • b = x • y • b) +instance additive.add_action [monoid α] [mul_action α β] : add_action (additive α) β := +{ vadd := (•) ∘ additive.to_mul, + zero_vadd := mul_action.one_smul, + add_vadd := mul_action.mul_smul } + +@[simp] lemma additive.of_mul_vadd [monoid α] [mul_action α β] (a : α) (b : β) : + additive.of_mul a +ᵥ b = a • b := +rfl + +instance multiplicative.mul_action [add_monoid α] [add_action α β] : + mul_action (multiplicative α) β := +{ smul := (+ᵥ) ∘ multiplicative.to_add, + one_smul := add_action.zero_vadd, + mul_smul := add_action.add_vadd } + +@[simp] lemma multiplicative.of_add_smul [add_monoid α] [add_action α β] (a : α) (b : β) : + multiplicative.of_add a • b = a +ᵥ b := +rfl + /-! ### (Pre)transitive action From 52a454af5e524961fcc26fb1bde541b76f4e0292 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Mon, 2 May 2022 13:45:00 +0000 Subject: [PATCH 005/153] feat(category_theory/limits): pushouts and pullbacks in the opposite category (#13495) This PR adds duality isomorphisms for the categories `wide_pushout_shape`, `wide_pullback_shape`, `walking_span`, `walking_cospan` and produce pullbacks/pushouts in the opposite category when pushouts/pullbacks exist. --- src/category_theory/limits/opposites.lean | 16 +++- .../limits/shapes/pullbacks.lean | 8 ++ .../limits/shapes/wide_pullbacks.lean | 76 ++++++++++++++++++- 3 files changed, 98 insertions(+), 2 deletions(-) diff --git a/src/category_theory/limits/opposites.lean b/src/category_theory/limits/opposites.lean index f93f9c6349383..e36b935129dfa 100644 --- a/src/category_theory/limits/opposites.lean +++ b/src/category_theory/limits/opposites.lean @@ -10,7 +10,7 @@ import category_theory.discrete_category # Limits in `C` give colimits in `Cᵒᵖ`. We also give special cases for (co)products, -but not yet for pullbacks / pushouts or for (co)equalizers. +(co)equalizers, and pullbacks / pushouts. -/ @@ -315,4 +315,18 @@ lemma has_finite_limits_opposite [has_finite_colimits C] : has_finite_limits Cᵒᵖ := { out := λ J 𝒟 𝒥, by { resetI, apply_instance, }, } +lemma has_pullbacks_opposite [has_pushouts C] : has_pullbacks Cᵒᵖ := +begin + haveI : has_colimits_of_shape walking_cospan.{v₁}ᵒᵖ C := + has_colimits_of_shape_of_equivalence walking_cospan_op_equiv.symm, + apply has_limits_of_shape_op_of_has_colimits_of_shape, +end + +lemma has_pushouts_opposite [has_pullbacks C] : has_pushouts Cᵒᵖ := +begin + haveI : has_limits_of_shape walking_span.{v₁}ᵒᵖ C := + has_limits_of_shape_of_equivalence walking_span_op_equiv.symm, + apply has_colimits_of_shape_op_of_has_limits_of_shape, +end + end category_theory.limits diff --git a/src/category_theory/limits/shapes/pullbacks.lean b/src/category_theory/limits/shapes/pullbacks.lean index 615ed787333fc..1f0e6bc85aad4 100644 --- a/src/category_theory/limits/shapes/pullbacks.lean +++ b/src/category_theory/limits/shapes/pullbacks.lean @@ -2157,4 +2157,12 @@ lemma has_pushouts_of_has_colimit_span has_pushouts C := { has_colimit := λ F, has_colimit_of_iso (diagram_iso_span F) } +/-- The duality equivalence `walking_spanᵒᵖ ≌ walking_cospan` -/ +def walking_span_op_equiv : walking_spanᵒᵖ ≌ walking_cospan := +wide_pushout_shape_op_equiv _ + +/-- The duality equivalence `walking_cospanᵒᵖ ≌ walking_span` -/ +def walking_cospan_op_equiv : walking_cospanᵒᵖ ≌ walking_span := +wide_pullback_shape_op_equiv _ + end category_theory.limits diff --git a/src/category_theory/limits/shapes/wide_pullbacks.lean b/src/category_theory/limits/shapes/wide_pullbacks.lean index 3c5202c0ff198..b2183b92546df 100644 --- a/src/category_theory/limits/shapes/wide_pullbacks.lean +++ b/src/category_theory/limits/shapes/wide_pullbacks.lean @@ -26,7 +26,7 @@ pullbacks and finite wide pullbacks. universes v u -open category_theory category_theory.limits +open category_theory category_theory.limits opposite namespace category_theory.limits @@ -371,4 +371,78 @@ end end wide_pushout +variable (J) + +/-- The action on morphisms of the obvious functor + `wide_pullback_shape_op : wide_pullback_shape J ⥤ (wide_pushout_shape J)ᵒᵖ`-/ +def wide_pullback_shape_op_map : Π (X Y : wide_pullback_shape J), + (X ⟶ Y) → ((op X : (wide_pushout_shape J)ᵒᵖ) ⟶ (op Y : (wide_pushout_shape J)ᵒᵖ)) +| _ _ (wide_pullback_shape.hom.id X) := quiver.hom.op (wide_pushout_shape.hom.id _) +| _ _ (wide_pullback_shape.hom.term j) := quiver.hom.op (wide_pushout_shape.hom.init _) + +/-- The obvious functor `wide_pullback_shape J ⥤ (wide_pushout_shape J)ᵒᵖ` -/ +@[simps] +def wide_pullback_shape_op : wide_pullback_shape J ⥤ (wide_pushout_shape J)ᵒᵖ := +{ obj := λ X, op X, + map := wide_pullback_shape_op_map J, } + +/-- The action on morphisms of the obvious functor +`wide_pushout_shape_op : `wide_pushout_shape J ⥤ (wide_pullback_shape J)ᵒᵖ` -/ +def wide_pushout_shape_op_map : Π (X Y : wide_pushout_shape J), + (X ⟶ Y) → ((op X : (wide_pullback_shape J)ᵒᵖ) ⟶ (op Y : (wide_pullback_shape J)ᵒᵖ)) +| _ _ (wide_pushout_shape.hom.id X) := quiver.hom.op (wide_pullback_shape.hom.id _) +| _ _ (wide_pushout_shape.hom.init j) := quiver.hom.op (wide_pullback_shape.hom.term _) + +/-- The obvious functor `wide_pushout_shape J ⥤ (wide_pullback_shape J)ᵒᵖ` -/ +@[simps] +def wide_pushout_shape_op : wide_pushout_shape J ⥤ (wide_pullback_shape J)ᵒᵖ := +{ obj := λ X, op X, + map := wide_pushout_shape_op_map J, } + +/-- The obvious functor `(wide_pullback_shape J)ᵒᵖ ⥤ wide_pushout_shape J`-/ +@[simps] +def wide_pullback_shape_unop : (wide_pullback_shape J)ᵒᵖ ⥤ wide_pushout_shape J := +(wide_pullback_shape_op J).left_op + +/-- The obvious functor `(wide_pushout_shape J)ᵒᵖ ⥤ wide_pullback_shape J` -/ +@[simps] +def wide_pushout_shape_unop : (wide_pushout_shape J)ᵒᵖ ⥤ wide_pullback_shape J := +(wide_pushout_shape_op J).left_op + +/-- The inverse of the unit isomorphism of the equivalence +`wide_pushout_shape_op_equiv : (wide_pushout_shape J)ᵒᵖ ≌ wide_pullback_shape J` -/ +def wide_pushout_shape_op_unop : wide_pushout_shape_unop J ⋙ wide_pullback_shape_op J ≅ 𝟭 _ := +nat_iso.of_components (λ X, iso.refl _) (λ X Y f, dec_trivial) + +/-- The counit isomorphism of the equivalence +`wide_pullback_shape_op_equiv : (wide_pullback_shape J)ᵒᵖ ≌ wide_pushout_shape J` -/ +def wide_pushout_shape_unop_op : wide_pushout_shape_op J ⋙ wide_pullback_shape_unop J ≅ 𝟭 _ := +nat_iso.of_components (λ X, iso.refl _) (λ X Y f, dec_trivial) + +/-- The inverse of the unit isomorphism of the equivalence +`wide_pullback_shape_op_equiv : (wide_pullback_shape J)ᵒᵖ ≌ wide_pushout_shape J` -/ +def wide_pullback_shape_op_unop : wide_pullback_shape_unop J ⋙ wide_pushout_shape_op J ≅ 𝟭 _ := +nat_iso.of_components (λ X, iso.refl _) (λ X Y f, dec_trivial) + +/-- The counit isomorphism of the equivalence +`wide_pushout_shape_op_equiv : (wide_pushout_shape J)ᵒᵖ ≌ wide_pullback_shape J` -/ +def wide_pullback_shape_unop_op : wide_pullback_shape_op J ⋙ wide_pushout_shape_unop J ≅ 𝟭 _ := +nat_iso.of_components (λ X, iso.refl _) (λ X Y f, dec_trivial) + +/-- The duality equivalence `(wide_pushout_shape J)ᵒᵖ ≌ wide_pullback_shape J` -/ +@[simps] +def wide_pushout_shape_op_equiv : (wide_pushout_shape J)ᵒᵖ ≌ wide_pullback_shape J := +{ functor := wide_pushout_shape_unop J, + inverse := wide_pullback_shape_op J, + unit_iso := (wide_pushout_shape_op_unop J).symm, + counit_iso := wide_pullback_shape_unop_op J, } + +/-- The duality equivalence `(wide_pullback_shape J)ᵒᵖ ≌ wide_pushout_shape J` -/ +@[simps] +def wide_pullback_shape_op_equiv : (wide_pullback_shape J)ᵒᵖ ≌ wide_pushout_shape J := +{ functor := wide_pullback_shape_unop J, + inverse := wide_pushout_shape_op J, + unit_iso := (wide_pullback_shape_op_unop J).symm, + counit_iso := wide_pushout_shape_unop_op J, } + end category_theory.limits From 3fde0827a555c4df0562c60317c3bbe3e0781bb5 Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Mon, 2 May 2022 13:45:01 +0000 Subject: [PATCH 006/153] refactor(topology/algebra/order): reorganize, simplify proofs (#13716) * Prove `has_compact_mul_support.is_compact_range` * Simplify the proof of `continuous.exists_forall_le_of_has_compact_mul_support` and `continuous.bdd_below_range_of_has_compact_mul_support` using `has_compact_mul_support.is_compact_range`. * Reorder `topology.algebra.order.basic` so that `is_compact.bdd_below` and friends are together with all results about `order_closed_topology`. * Move `continuous.bdd_below_range_of_has_compact_mul_support` (and dual) to `topology.algebra.order.basic` --- src/algebra/support.lean | 10 ++++ src/topology/algebra/order/basic.lean | 73 ++++++++++++++----------- src/topology/algebra/order/compact.lean | 38 +++---------- src/topology/support.lean | 19 +++++++ 4 files changed, 80 insertions(+), 60 deletions(-) diff --git a/src/algebra/support.lean b/src/algebra/support.lean index 6d4e88c9b9d51..d900375d0d67d 100644 --- a/src/algebra/support.lean +++ b/src/algebra/support.lean @@ -71,6 +71,16 @@ by { simp_rw [← subset_empty_iff, mul_support_subset_iff', funext_iff], simp } (mul_support f).nonempty ↔ f ≠ 1 := by rw [← ne_empty_iff_nonempty, ne.def, mul_support_eq_empty_iff] +@[to_additive] +lemma range_subset_insert_image_mul_support (f : α → M) : + range f ⊆ insert 1 (f '' mul_support f) := +begin + intros y hy, + rcases eq_or_ne y 1 with rfl|h2y, + { exact mem_insert _ _ }, + { obtain ⟨x, rfl⟩ := hy, refine mem_insert_of_mem _ ⟨x, h2y, rfl⟩ } +end + @[simp, to_additive] lemma mul_support_one' : mul_support (1 : α → M) = ∅ := mul_support_eq_empty_iff.2 rfl diff --git a/src/topology/algebra/order/basic.lean b/src/topology/algebra/order/basic.lean index df9d85550cfca..d90d7a85bcb9c 100644 --- a/src/topology/algebra/order/basic.lean +++ b/src/topology/algebra/order/basic.lean @@ -629,6 +629,48 @@ lemma dense.exists_between [densely_ordered α] {s : set α} (hs : dense s) {x y ∃ z ∈ s, z ∈ Ioo x y := hs.exists_mem_open is_open_Ioo (nonempty_Ioo.2 h) +variables [nonempty α] [topological_space β] + +/-- A compact set is bounded below -/ +lemma is_compact.bdd_below {s : set α} (hs : is_compact s) : bdd_below s := +begin + by_contra H, + rcases hs.elim_finite_subcover_image (λ x (_ : x ∈ s), @is_open_Ioi _ _ _ _ x) _ + with ⟨t, st, ft, ht⟩, + { refine H (ft.bdd_below.imp $ λ C hC y hy, _), + rcases mem_Union₂.1 (ht hy) with ⟨x, hx, xy⟩, + exact le_trans (hC hx) (le_of_lt xy) }, + { refine λ x hx, mem_Union₂.2 (not_imp_comm.1 _ H), + exact λ h, ⟨x, λ y hy, le_of_not_lt (h.imp $ λ ys, ⟨_, hy, ys⟩)⟩ } +end + +/-- A compact set is bounded above -/ +lemma is_compact.bdd_above {s : set α} (hs : is_compact s) : bdd_above s := +@is_compact.bdd_below (order_dual α) _ _ _ _ _ hs + +/-- A continuous function is bounded below on a compact set. -/ +lemma is_compact.bdd_below_image {f : β → α} {K : set β} + (hK : is_compact K) (hf : continuous_on f K) : bdd_below (f '' K) := +(hK.image_of_continuous_on hf).bdd_below + +/-- A continuous function is bounded above on a compact set. -/ +lemma is_compact.bdd_above_image {f : β → α} {K : set β} + (hK : is_compact K) (hf : continuous_on f K) : bdd_above (f '' K) := +@is_compact.bdd_below_image (order_dual α) _ _ _ _ _ _ _ _ hK hf + +/-- A continuous function with compact support is bounded below. -/ +@[to_additive /-" A continuous function with compact support is bounded below. "-/] +lemma continuous.bdd_below_range_of_has_compact_mul_support [has_one α] {f : β → α} + (hf : continuous f) (h : has_compact_mul_support f) : bdd_below (range f) := +(h.is_compact_range hf).bdd_below + +/-- A continuous function with compact support is bounded above. -/ +@[to_additive /-" A continuous function with compact support is bounded above. "-/] +lemma continuous.bdd_above_range_of_has_compact_mul_support [has_one α] + {f : β → α} (hf : continuous f) (h : has_compact_mul_support f) : + bdd_above (range f) := +@continuous.bdd_below_range_of_has_compact_mul_support (order_dual α) _ _ _ _ _ _ _ _ hf h + end linear_order end order_closed_topology @@ -2184,37 +2226,6 @@ lemma exists_seq_tendsto_Inf {α : Type*} [conditionally_complete_linear_order ∃ (u : ℕ → α), antitone u ∧ tendsto u at_top (𝓝 (Inf S)) ∧ (∀ n, u n ∈ S) := @exists_seq_tendsto_Sup (order_dual α) _ _ _ _ S hS hS' -/-- A compact set is bounded below -/ -lemma is_compact.bdd_below {α : Type u} [topological_space α] [linear_order α] - [order_closed_topology α] [nonempty α] {s : set α} (hs : is_compact s) : bdd_below s := -begin - by_contra H, - rcases hs.elim_finite_subcover_image (λ x (_ : x ∈ s), @is_open_Ioi _ _ _ _ x) _ - with ⟨t, st, ft, ht⟩, - { refine H (ft.bdd_below.imp $ λ C hC y hy, _), - rcases mem_Union₂.1 (ht hy) with ⟨x, hx, xy⟩, - exact le_trans (hC hx) (le_of_lt xy) }, - { refine λ x hx, mem_Union₂.2 (not_imp_comm.1 _ H), - exact λ h, ⟨x, λ y hy, le_of_not_lt (h.imp $ λ ys, ⟨_, hy, ys⟩)⟩ } -end - -/-- A compact set is bounded above -/ -lemma is_compact.bdd_above {α : Type u} [topological_space α] [linear_order α] - [order_closed_topology α] : Π [nonempty α] {s : set α}, is_compact s → bdd_above s := -@is_compact.bdd_below (order_dual α) _ _ _ - -/-- A continuous function is bounded below on a compact set. -/ -lemma is_compact.bdd_below_image {α : Type u} [topological_space α] [linear_order α] - [order_closed_topology α] [nonempty α] [topological_space γ] {f : γ → α} {K : set γ} - (hK : is_compact K) (hf : continuous_on f K) : bdd_below (f '' K) := -(hK.image_of_continuous_on hf).bdd_below - -/-- A continuous function is bounded above on a compact set. -/ -lemma is_compact.bdd_above_image {α : Type u} [topological_space α] [linear_order α] - [order_closed_topology α] [nonempty α] [topological_space γ] {f : γ → α} {K : set γ} - (hK : is_compact K) (hf : continuous_on f K) : bdd_above (f '' K) := -@is_compact.bdd_below_image _ (order_dual α) _ _ _ _ _ _ _ hK hf - end order_topology section densely_ordered diff --git a/src/topology/algebra/order/compact.lean b/src/topology/algebra/order/compact.lean index d6cb0371f26e0..995c7f12478dd 100644 --- a/src/topology/algebra/order/compact.lean +++ b/src/topology/algebra/order/compact.lean @@ -9,9 +9,11 @@ import topology.algebra.order.intermediate_value # Compactness of a closed interval In this file we prove that a closed interval in a conditionally complete linear ordered type with -order topology (or a product of such types) is compact. We also prove the extreme value theorem -(`is_compact.exists_forall_le`, `is_compact.exists_forall_ge`): a continuous function on a compact -set takes its minimum and maximum values. +order topology (or a product of such types) is compact. + +We prove the extreme value theorem (`is_compact.exists_forall_le`, `is_compact.exists_forall_ge`): +a continuous function on a compact set takes its minimum and maximum values. We provide many +variations of this theorem. We also prove that the image of a closed interval under a continuous map is a closed interval, see `continuous_on.image_Icc`. @@ -195,7 +197,7 @@ let ⟨x, hxs, hx, _⟩ := hs.exists_Inf_image_eq_and_le ne_s hf in ⟨x, hxs, h lemma is_compact.exists_Sup_image_eq : ∀ {s : set β}, is_compact s → s.nonempty → ∀ {f : β → α}, continuous_on f s → - ∃ x ∈ s, Sup (f '' s) = f x := + ∃ x ∈ s, Sup (f '' s) = f x := @is_compact.exists_Inf_image_eq (order_dual α) _ _ _ _ _ lemma eq_Icc_of_connected_compact {s : set α} (h₁ : is_connected s) (h₂ : is_compact s) : @@ -295,12 +297,9 @@ lemma _root_.continuous.exists_forall_le_of_has_compact_mul_support [nonempty β {f : β → α} (hf : continuous f) (h : has_compact_mul_support f) : ∃ (x : β), ∀ (y : β), f x ≤ f y := begin - -- Proof sketch: we use `continuous.exists_forall_le'` with as `x₀` any element outside the - -- support of `f`, if such an element exists (and otherwise an arbitrary element). - refine hf.exists_forall_le' (classical.epsilon $ λ x, f x = 1) - (eventually_of_mem h.compl_mem_cocompact $ λ x hx, _), - have : f x = 1 := nmem_mul_support.mp (mt (λ h2x, subset_closure h2x) hx), - exact ((classical.epsilon_spec ⟨x, this⟩).trans this.symm).le + obtain ⟨_, ⟨x, rfl⟩, hx⟩ := (h.is_compact_range hf).exists_is_least (range_nonempty _), + rw [mem_lower_bounds, forall_range_iff] at hx, + exact ⟨x, hx⟩, end /-- A continuous function with compact support has a global maximum. -/ @@ -310,25 +309,6 @@ lemma continuous.exists_forall_ge_of_has_compact_mul_support [nonempty β] [has_ ∃ (x : β), ∀ (y : β), f y ≤ f x := @continuous.exists_forall_le_of_has_compact_mul_support (order_dual α) _ _ _ _ _ _ _ _ hf h -/-- A continuous function with compact support is bounded below. -/ -@[to_additive] -lemma continuous.bdd_below_range_of_has_compact_mul_support [has_one α] - {f : β → α} (hf : continuous f) (h : has_compact_mul_support f) : - bdd_below (range f) := -begin - casesI is_empty_or_nonempty β with hβ hβ, - { rw range_eq_empty_iff.mpr, exact bdd_below_empty, exact hβ }, - obtain ⟨x, hx⟩ := hf.exists_forall_le_of_has_compact_mul_support h, - refine ⟨f x, _⟩, rintro _ ⟨x', rfl⟩, exact hx x' -end - -/-- A continuous function with compact support is bounded above. -/ -@[to_additive] -lemma continuous.bdd_above_range_of_has_compact_mul_support [has_one α] - {f : β → α} (hf : continuous f) (h : has_compact_mul_support f) : - bdd_above (range f) := -@continuous.bdd_below_range_of_has_compact_mul_support (order_dual α) _ _ _ _ _ _ _ hf h - lemma is_compact.continuous_Sup {f : γ → β → α} {K : set β} (hK : is_compact K) (hf : continuous ↿f) : continuous (λ x, Sup (f x '' K)) := diff --git a/src/topology/support.lean b/src/topology/support.lean index 80c8595fc5390..708e55ebc648f 100644 --- a/src/topology/support.lean +++ b/src/topology/support.lean @@ -59,6 +59,17 @@ by rw [mul_tsupport, closure_empty_iff, mul_support_eq_empty_iff] lemma image_eq_zero_of_nmem_mul_tsupport {f : X → α} {x : X} (hx : x ∉ mul_tsupport f) : f x = 1 := mul_support_subset_iff'.mp (subset_mul_tsupport f) x hx +@[to_additive] +lemma range_subset_insert_image_mul_tsupport (f : X → α) : + range f ⊆ insert 1 (f '' mul_tsupport f) := +(range_subset_insert_image_mul_support f).trans $ + insert_subset_insert $ image_subset _ subset_closure + +@[to_additive] +lemma range_eq_image_mul_tsupport_or (f : X → α) : + range f = f '' mul_tsupport f ∨ range f = insert 1 (f '' mul_tsupport f) := +(wcovby_insert _ _).eq_or_eq (image_subset_range _ _) (range_subset_insert_image_mul_tsupport f) + end one section @@ -111,6 +122,14 @@ lemma has_compact_mul_support_iff_eventually_eq : λ h, let ⟨C, hC⟩ := mem_coclosed_compact'.mp h in compact_of_is_closed_subset hC.2.1 (is_closed_mul_tsupport _) (closure_minimal hC.2.2 hC.1)⟩ +@[to_additive] +lemma has_compact_mul_support.is_compact_range [topological_space β] + (h : has_compact_mul_support f) (hf : continuous f) : is_compact (range f) := +begin + cases range_eq_image_mul_tsupport_or f with h2 h2; rw [h2], + exacts [h.image hf, (h.image hf).insert 1] +end + @[to_additive] lemma has_compact_mul_support.mono' {f' : α → γ} (hf : has_compact_mul_support f) (hff' : mul_support f' ⊆ mul_tsupport f) : has_compact_mul_support f' := From b063c2861e09afb51c6eb0925319bd8cce1e79a7 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 2 May 2022 13:45:02 +0000 Subject: [PATCH 007/153] =?UTF-8?q?fix(src/tactic/alias):=20Support=20`ali?= =?UTF-8?q?as=20foo=20=E2=86=94=20..`=20as=20documented=20(#13743)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit the current code and the single(!) use of this feature work only if you write `alias foo ↔ . .` which is very odd. --- src/data/polynomial/degree/definitions.lean | 2 +- src/tactic/alias.lean | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/src/data/polynomial/degree/definitions.lean b/src/data/polynomial/degree/definitions.lean index 7fdca44f69ed9..60e1e96c701cf 100644 --- a/src/data/polynomial/degree/definitions.lean +++ b/src/data/polynomial/degree/definitions.lean @@ -155,7 +155,7 @@ lemma nat_degree_lt_iff_degree_lt (hp : p ≠ 0) : p.nat_degree < n ↔ p.degree < ↑n := with_bot.get_or_else_bot_lt_iff $ degree_eq_bot.not.mpr hp -alias polynomial.nat_degree_le_iff_degree_le ↔ . . +alias polynomial.nat_degree_le_iff_degree_le ↔ .. lemma nat_degree_le_nat_degree [semiring S] {q : S[X]} (hpq : p.degree ≤ q.degree) : p.nat_degree ≤ q.nat_degree := diff --git a/src/tactic/alias.lean b/src/tactic/alias.lean index bc09269102fff..1353c9fac8798 100644 --- a/src/tactic/alias.lean +++ b/src/tactic/alias.lean @@ -138,7 +138,7 @@ do old ← ident, do { tk "↔" <|> tk "<->", (left, right) ← - mcond ((tk "." *> tk "." >> pure tt) <|> pure ff) + 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) left `iff.mp, From 4113e00adf60e50fca319b741c5a6a77ecfe63b5 Mon Sep 17 00:00:00 2001 From: Oliver Nash Date: Mon, 2 May 2022 13:45:03 +0000 Subject: [PATCH 008/153] feat(linear_algebra/affine_space/basis): add lemma `affine_basis.linear_combination_coord_eq_self` (#13763) Formalized as part of the Sphere Eversion project. --- src/linear_algebra/affine_space/basis.lean | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/src/linear_algebra/affine_space/basis.lean b/src/linear_algebra/affine_space/basis.lean index 88682c2f1b4f5..20402c7e4bc00 100644 --- a/src/linear_algebra/affine_space/basis.lean +++ b/src/linear_algebra/affine_space/basis.lean @@ -151,6 +151,15 @@ begin exact b.coord_apply_combination_of_mem (finset.mem_univ i) hw, end +/-- A variant of `affine_basis.affine_combination_coord_eq_self` for the special case when the +affine space is a module so we can talk about linear combinations. -/ +@[simp] lemma linear_combination_coord_eq_self [fintype ι] (b : affine_basis ι k V) (v : V) : + ∑ i, (b.coord i v) • (b.points i) = v := +begin + have hb := b.affine_combination_coord_eq_self v, + rwa finset.univ.affine_combination_eq_linear_combination _ _ (b.sum_coord_apply_eq_one v) at hb, +end + lemma ext_elem [fintype ι] {q₁ q₂ : P} (h : ∀ i, b.coord i q₁ = b.coord i q₂) : q₁ = q₂ := begin rw [← b.affine_combination_coord_eq_self q₁, ← b.affine_combination_coord_eq_self q₂], From cf5fa84cb487d04e580d45ba52839e91cb0fe13e Mon Sep 17 00:00:00 2001 From: Oliver Nash Date: Mon, 2 May 2022 13:45:04 +0000 Subject: [PATCH 009/153] feat(analysis/normed_space/add_torsor_bases): add lemma `smooth_barycentric_coord` (#13764) Formalized as part of the Sphere Eversion project. --- src/analysis/normed_space/add_torsor_bases.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/src/analysis/normed_space/add_torsor_bases.lean b/src/analysis/normed_space/add_torsor_bases.lean index 3a4dc2b213f0d..c8a8e98bf05aa 100644 --- a/src/analysis/normed_space/add_torsor_bases.lean +++ b/src/analysis/normed_space/add_torsor_bases.lean @@ -5,6 +5,7 @@ Authors: Oliver Nash -/ import analysis.normed_space.banach import analysis.normed_space.finite_dimension +import analysis.calculus.affine_map import analysis.convex.combination import linear_algebra.affine_space.basis import linear_algebra.affine_space.finite_dimensional @@ -40,6 +41,9 @@ lemma is_open_map_barycentric_coord [nontrivial ι] (i : ι) : is_open_map (b.coord i) := (b.coord i).is_open_map (continuous_barycentric_coord b i) (b.surjective_coord i) +lemma smooth_barycentric_coord (b : affine_basis ι 𝕜 E) (i : ι) : cont_diff 𝕜 ⊤ (b.coord i) := +(⟨b.coord i, continuous_barycentric_coord b i⟩ : E →A[𝕜] 𝕜).cont_diff + end barycentric open set From 4fe734da35688d84ba11b786d7f8af26ab5cdb53 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 2 May 2022 13:45:06 +0000 Subject: [PATCH 010/153] fix(algebra/indicator_function): add missing decidable instances to lemma statements (#13834) This keeps the definition of `set.indicator` as non-computable, but ensures that when lemmas are applied they generalize to any decidable instances. --- src/algebra/indicator_function.lean | 63 +++++++++++++------ .../simple_graph/inc_matrix.lean | 6 +- src/geometry/euclidean/circumcenter.lean | 2 +- .../function/conditional_expectation.lean | 1 + .../function/uniform_integrable.lean | 2 +- src/measure_theory/integral/bochner.lean | 24 ++++--- .../probability_mass_function/monad.lean | 2 +- src/topology/instances/ennreal.lean | 7 +-- 8 files changed, 67 insertions(+), 40 deletions(-) diff --git a/src/algebra/indicator_function.lean b/src/algebra/indicator_function.lean index 25bde364046b8..2dffdf02a2538 100644 --- a/src/algebra/indicator_function.lean +++ b/src/algebra/indicator_function.lean @@ -21,12 +21,14 @@ But since it is usually used to restrict a function to a certain set `s`, we let the indicator function take the value `f x` for some function `f`, instead of `1`. If the usual indicator function is needed, just set `f` to be the constant function `λx, 1`. +The indicator function is implemented non-computably, to avoid having to pass around `decidable` +arguments. This is in contrast with the design of `pi.single` or `set.piecewise`. + ## Tags indicator, characteristic -/ -noncomputable theory -open_locale classical big_operators +open_locale big_operators open function variables {α β ι M N : Type*} @@ -37,30 +39,40 @@ section has_one variables [has_one M] [has_one N] {s t : set α} {f g : α → M} {a : α} /-- `indicator s f a` is `f a` if `a ∈ s`, `0` otherwise. -/ -def indicator {M} [has_zero M] (s : set α) (f : α → M) : α → M := λ x, if x ∈ s then f x else 0 +noncomputable def indicator {M} [has_zero M] (s : set α) (f : α → M) : α → M +| x := by haveI := classical.dec_pred (∈ s); exact if x ∈ s then f x else 0 /-- `mul_indicator s f a` is `f a` if `a ∈ s`, `1` otherwise. -/ @[to_additive] -def mul_indicator (s : set α) (f : α → M) : α → M := λ x, if x ∈ s then f x else 1 +noncomputable def mul_indicator (s : set α) (f : α → M) : α → M +| x := by haveI := classical.dec_pred (∈ s); exact if x ∈ s then f x else 1 -@[simp, to_additive] lemma piecewise_eq_mul_indicator : s.piecewise f 1 = s.mul_indicator f := rfl +@[simp, to_additive] lemma piecewise_eq_mul_indicator [decidable_pred (∈ s)] : + s.piecewise f 1 = s.mul_indicator f := +funext $ λ x, @if_congr _ _ _ _ (id _) _ _ _ _ iff.rfl rfl rfl -@[to_additive] lemma mul_indicator_apply (s : set α) (f : α → M) (a : α) : - mul_indicator s f a = if a ∈ s then f a else 1 := rfl +@[to_additive] lemma mul_indicator_apply (s : set α) (f : α → M) (a : α) [decidable (a ∈ s)] : + mul_indicator s f a = if a ∈ s then f a else 1 := by convert rfl @[simp, to_additive] lemma mul_indicator_of_mem (h : a ∈ s) (f : α → M) : - mul_indicator s f a = f a := if_pos h + mul_indicator s f a = f a := +by { letI := classical.dec (a ∈ s), exact if_pos h } @[simp, to_additive] lemma mul_indicator_of_not_mem (h : a ∉ s) (f : α → M) : - mul_indicator s f a = 1 := if_neg h + mul_indicator s f a = 1 := +by { letI := classical.dec (a ∈ s), exact if_neg h } @[to_additive] lemma mul_indicator_eq_one_or_self (s : set α) (f : α → M) (a : α) : mul_indicator s f a = 1 ∨ mul_indicator s f a = f a := -if h : a ∈ s then or.inr (mul_indicator_of_mem h f) else or.inl (mul_indicator_of_not_mem h f) +begin + by_cases h : a ∈ s, + { exact or.inr (mul_indicator_of_mem h f) }, + { exact or.inl (mul_indicator_of_not_mem h f) } +end @[simp, to_additive] lemma mul_indicator_apply_eq_self : s.mul_indicator f a = f a ↔ (a ∉ s → f a = 1) := -ite_eq_left_iff.trans $ by rw [@eq_comm _ (f a)] +by letI := classical.dec (a ∈ s); exact ite_eq_left_iff.trans (by rw [@eq_comm _ (f a)]) @[simp, to_additive] lemma mul_indicator_eq_self : s.mul_indicator f = f ↔ mul_support f ⊆ s := by simp only [funext_iff, subset_def, mem_mul_support, mul_indicator_apply_eq_self, not_imp_comm] @@ -71,7 +83,7 @@ by { rw mul_indicator_eq_self at h1 ⊢, exact subset.trans h1 h2 } @[simp, to_additive] lemma mul_indicator_apply_eq_one : mul_indicator s f a = 1 ↔ (a ∈ s → f a = 1) := -ite_eq_right_iff +by letI := classical.dec (a ∈ s); exact ite_eq_right_iff @[simp, to_additive] lemma mul_indicator_eq_one : mul_indicator s f = (λ x, 1) ↔ disjoint (mul_support f) s := @@ -108,7 +120,7 @@ mul_indicator_eq_self.2 subset.rfl @[simp, to_additive] lemma mul_indicator_range_comp {ι : Sort*} (f : ι → α) (g : α → M) : mul_indicator (range f) g ∘ f = g ∘ f := -piecewise_range_comp _ _ _ +by letI := classical.dec_pred (∈ range f); exact piecewise_range_comp _ _ _ @[to_additive] lemma mul_indicator_congr (h : eq_on f g s) : mul_indicator s f = mul_indicator s g := @@ -142,9 +154,10 @@ funext $ λx, by { simp only [mul_indicator], split_ifs, repeat {simp * at * {co mul_indicator (s ∩ mul_support f) f = mul_indicator s f := by rw [← mul_indicator_mul_indicator, mul_indicator_mul_support] -@[to_additive] lemma comp_mul_indicator (h : M → β) (f : α → M) {s : set α} {x : α} : +@[to_additive] lemma comp_mul_indicator (h : M → β) (f : α → M) {s : set α} {x : α} + [decidable_pred (∈ s)] : h (s.mul_indicator f x) = s.piecewise (h ∘ f) (const α (h 1)) x := -s.apply_piecewise _ _ (λ _, h) +by letI := classical.dec_pred (∈ s); convert s.apply_piecewise f (const α 1) (λ _, h) @[to_additive] lemma mul_indicator_comp_right {s : set α} (f : β → α) {g : α → M} {x : β} : mul_indicator (f ⁻¹' s) (g ∘ f) x = mul_indicator s g (f x) := @@ -168,7 +181,7 @@ end @[to_additive] lemma mul_indicator_preimage (s : set α) (f : α → M) (B : set M) : (mul_indicator s f)⁻¹' B = s.ite (f ⁻¹' B) (1 ⁻¹' B) := -piecewise_preimage s f 1 B +by letI := classical.dec_pred (∈ s); exact piecewise_preimage s f 1 B @[to_additive] lemma mul_indicator_preimage_of_not_mem (s : set α) (f : α → M) {t : set M} (ht : (1:M) ∉ t) : @@ -264,7 +277,7 @@ end /-- `set.mul_indicator` as a `monoid_hom`. -/ @[to_additive "`set.indicator` as an `add_monoid_hom`."] -def mul_indicator_hom {α} (M) [mul_one_class M] (s : set α) : (α → M) →* (α → M) := +noncomputable def mul_indicator_hom {α} (M) [mul_one_class M] (s : set α) : (α → M) →* (α → M) := { to_fun := mul_indicator s, map_one' := mul_indicator_one M s, map_mul' := mul_indicator_mul s } @@ -374,7 +387,7 @@ lemma prod_mul_indicator_subset (f : α → M) {s t : finset α} (h : s ⊆ t) : prod_mul_indicator_subset_of_eq_one _ (λ a b, b) h (λ _, rfl) @[to_additive] lemma _root_.finset.prod_mul_indicator_eq_prod_filter - (s : finset ι) (f : ι → α → M) (t : ι → set α) (g : ι → α) : + (s : finset ι) (f : ι → α → M) (t : ι → set α) (g : ι → α) [decidable_pred (λ i, g i ∈ t i)]: ∏ i in s, mul_indicator (t i) (f i) (g i) = ∏ i in s.filter (λ i, g i ∈ t i), f i (g i) := begin refine (finset.prod_filter_mul_prod_filter_not s (λ i, g i ∈ t i) _).symm.trans _, @@ -392,6 +405,7 @@ end (s : ι → set α) {f : α → M} : (∀ (i ∈ I) (j ∈ I), i ≠ j → disjoint (s i) (s j)) → mul_indicator (⋃ i ∈ I, s i) f = λ a, ∏ i in I, mul_indicator (s i) f a := begin + classical, refine finset.induction_on I _ _, { intro h, funext, simp }, assume a I haI ih hI, @@ -439,7 +453,11 @@ funext (λ _, by simpa only [← inter_indicator_mul, pi.mul_apply, pi.one_apply lemma indicator_prod_one {s : set α} {t : set β} {x : α} {y : β} : (s ×ˢ t : set _).indicator (1 : _ → M) (x, y) = s.indicator 1 x * t.indicator 1 y := -by simp [indicator, ← ite_and] +begin + letI := classical.dec_pred (∈ s), + letI := classical.dec_pred (∈ t), + simp [indicator_apply, ← ite_and], +end end mul_zero_one_class @@ -452,7 +470,11 @@ variables [has_le M] @[to_additive] lemma mul_indicator_apply_le' (hfg : a ∈ s → f a ≤ y) (hg : a ∉ s → 1 ≤ y) : mul_indicator s f a ≤ y := -if ha : a ∈ s then by simpa [ha] using hfg ha else by simpa [ha] using hg ha +begin + by_cases ha : a ∈ s, + { simpa [ha] using hfg ha }, + { simpa [ha] using hg ha }, +end @[to_additive] lemma mul_indicator_le' (hfg : ∀ a ∈ s, f a ≤ g a) (hg : ∀ a ∉ s, 1 ≤ g a) : mul_indicator s f ≤ g := @@ -540,6 +562,7 @@ lemma indicator_le_indicator_nonneg {β} [linear_order β] [has_zero β] (s : se s.indicator f ≤ {x | 0 ≤ f x}.indicator f := begin intro x, + classical, simp_rw indicator_apply, split_ifs, { exact le_rfl, }, diff --git a/src/combinatorics/simple_graph/inc_matrix.lean b/src/combinatorics/simple_graph/inc_matrix.lean index e2bb5180723c3..44cea122c62a5 100644 --- a/src/combinatorics/simple_graph/inc_matrix.lean +++ b/src/combinatorics/simple_graph/inc_matrix.lean @@ -40,8 +40,6 @@ incidence matrix for each `simple_graph α` has the same type. arbitrary orientation of a simple graph. -/ -noncomputable theory - open finset matrix simple_graph sym2 open_locale big_operators matrix @@ -50,7 +48,7 @@ variables (R : Type*) {α : Type*} (G : simple_graph α) /-- `G.inc_matrix R` is the `α × sym2 α` matrix whose `(a, e)`-entry is `1` if `e` is incident to `a` and `0` otherwise. -/ -def inc_matrix [has_zero R] [has_one R] : matrix α (sym2 α) R := +noncomputable def inc_matrix [has_zero R] [has_one R] : matrix α (sym2 α) R := λ a, (G.incidence_set a).indicator 1 variables {R} @@ -70,6 +68,7 @@ variables [mul_zero_one_class R] {a b : α} {e : sym2 α} lemma inc_matrix_apply_mul_inc_matrix_apply : G.inc_matrix R a e * G.inc_matrix R b e = (G.incidence_set a ∩ G.incidence_set b).indicator 1 e := begin + classical, simp only [inc_matrix, set.indicator_apply, ←ite_and_mul_zero, pi.one_apply, mul_one, set.mem_inter_eq], congr, @@ -162,6 +161,7 @@ variables [fintype (sym2 α)] [semiring R] {a b : α} {e : sym2 α} lemma inc_matrix_mul_transpose_apply_of_adj (h : G.adj a b) : (G.inc_matrix R ⬝ (G.inc_matrix R)ᵀ) a b = (1 : R) := begin + classical, simp_rw [matrix.mul_apply, matrix.transpose_apply, inc_matrix_apply_mul_inc_matrix_apply, set.indicator_apply, pi.one_apply, sum_boole], convert nat.cast_one, diff --git a/src/geometry/euclidean/circumcenter.lean b/src/geometry/euclidean/circumcenter.lean index 72dfcf606f9df..bb4fbc109bac0 100644 --- a/src/geometry/euclidean/circumcenter.lean +++ b/src/geometry/euclidean/circumcenter.lean @@ -615,7 +615,7 @@ begin fs.subset_univ (λ i, zero_smul ℝ _), set.indicator_apply], - congr, funext, congr' 2 + congr, end omit V diff --git a/src/measure_theory/function/conditional_expectation.lean b/src/measure_theory/function/conditional_expectation.lean index 923c87cb640bb..593f0e8198d79 100644 --- a/src/measure_theory/function/conditional_expectation.lean +++ b/src/measure_theory/function/conditional_expectation.lean @@ -853,6 +853,7 @@ begin { refine lintegral_congr_ae (ae_restrict_of_ae _), refine (@indicator_const_Lp_coe_fn _ _ _ 2 _ _ _ hs hμs (1 : ℝ)).mono (λ x hx, _), rw hx, + classical, simp_rw set.indicator_apply, split_ifs; simp, }, rw [h_eq, lintegral_indicator _ hs, lintegral_const, measure.restrict_restrict hs], diff --git a/src/measure_theory/function/uniform_integrable.lean b/src/measure_theory/function/uniform_integrable.lean index d66f216b6c883..0821464f1df43 100644 --- a/src/measure_theory/function/uniform_integrable.lean +++ b/src/measure_theory/function/uniform_integrable.lean @@ -809,7 +809,7 @@ begin ((hf i).nnnorm.measurable_set_lt strongly_measurable_const))) (strongly_measurable.ae_strongly_measurable ((hf i).indicator (strongly_measurable_const.measurable_set_le (hf i).nnnorm))) hp), - { rw [indicator, pi.add_apply], + { rw [pi.add_apply, indicator_apply], split_ifs with hx, { rw [indicator_of_not_mem, add_zero], simpa using hx }, diff --git a/src/measure_theory/integral/bochner.lean b/src/measure_theory/integral/bochner.lean index 5c7bebc92e623..1347b71d05914 100644 --- a/src/measure_theory/integral/bochner.lean +++ b/src/measure_theory/integral/bochner.lean @@ -142,7 +142,7 @@ Bochner integral, simple function, function space, Lebesgue dominated convergenc -/ noncomputable theory -open_locale classical topological_space big_operators nnreal ennreal measure_theory +open_locale topological_space big_operators nnreal ennreal measure_theory open set filter topological_space ennreal emetric namespace measure_theory @@ -294,12 +294,13 @@ lemma integral_eq {m : measurable_space α} (μ : measure α) (f : α →ₛ F) f.integral μ = ∑ x in f.range, (μ (f ⁻¹' {x})).to_real • x := by simp [integral, set_to_simple_func, weighted_smul_apply] -lemma integral_eq_sum_filter {m : measurable_space α} (f : α →ₛ F) (μ : measure α) : +lemma integral_eq_sum_filter [decidable_pred (λ x : F, x ≠ 0)] {m : measurable_space α} (f : α →ₛ F) + (μ : measure α) : f.integral μ = ∑ x in f.range.filter (λ x, x ≠ 0), (μ (f ⁻¹' {x})).to_real • x := -by { rw [integral_def, set_to_simple_func_eq_sum_filter], simp_rw weighted_smul_apply, } +by { rw [integral_def, set_to_simple_func_eq_sum_filter], simp_rw weighted_smul_apply, congr } /-- The Bochner integral is equal to a sum over any set that includes `f.range` (except `0`). -/ -lemma integral_eq_sum_of_subset {f : α →ₛ F} {s : finset F} +lemma integral_eq_sum_of_subset [decidable_pred (λ x : F, x ≠ 0)] {f : α →ₛ F} {s : finset F} (hs : f.range.filter (λ x, x ≠ 0) ⊆ s) : f.integral μ = ∑ x in s, (μ (f ⁻¹' {x})).to_real • x := begin rw [simple_func.integral_eq_sum_filter, finset.sum_subset hs], @@ -310,7 +311,7 @@ end @[simp] lemma integral_const {m : measurable_space α} (μ : measure α) (y : F) : (const α y).integral μ = (μ univ).to_real • y := -calc (const α y).integral μ = ∑ z in {y}, (μ ((const α y) ⁻¹' {z})).to_real • z : +by classical; calc (const α y).integral μ = ∑ z in {y}, (μ ((const α y) ⁻¹' {z})).to_real • z : integral_eq_sum_of_subset $ (filter_subset _ _).trans (range_const_subset _ _) ... = (μ univ).to_real • y : by simp @@ -318,6 +319,7 @@ calc (const α y).integral μ = ∑ z in {y}, (μ ((const α y) ⁻¹' {z})).to_ {s : set α} (hs : measurable_set s) : (piecewise s hs f 0).integral μ = f.integral (μ.restrict s) := begin + classical, refine (integral_eq_sum_of_subset _).trans ((sum_congr rfl $ λ y hy, _).trans (integral_eq_sum_filter _ _).symm), { intros y hy, @@ -326,7 +328,8 @@ begin rcases hy with ⟨⟨rfl, -⟩|⟨x, hxs, rfl⟩, h₀⟩, exacts [(h₀ rfl).elim, ⟨set.mem_range_self _, h₀⟩] }, { dsimp, - rw [indicator_preimage_of_not_mem, measure.restrict_apply (f.measurable_set_preimage _)], + rw [set.piecewise_eq_indicator, indicator_preimage_of_not_mem, + measure.restrict_apply (f.measurable_set_preimage _)], exact λ h₀, (mem_filter.1 hy).2 (eq.symm h₀) } end @@ -687,10 +690,15 @@ variables [normed_group E] [normed_space ℝ E] [complete_space E] [nondiscrete_normed_field 𝕜] [normed_space 𝕜 E] [smul_comm_class ℝ 𝕜 E] [normed_group F] [normed_space ℝ F] [complete_space F] +section +open_locale classical + /-- The Bochner integral -/ def integral {m : measurable_space α} (μ : measure α) (f : α → E) : E := if hf : integrable f μ then L1.integral (hf.to_L1 f) else 0 +end + /-! In the notation for integrals, an expression like `∫ x, g ∥x∥ ∂μ` will not be parsed correctly, and needs parentheses. We do not set the binding power of `r` to `0`, because then `∫ x, f x = 0` will be parsed incorrectly. -/ @@ -707,7 +715,7 @@ variables {f g : α → E} {m : measurable_space α} {μ : measure α} lemma integral_eq (f : α → E) (hf : integrable f μ) : ∫ a, f a ∂μ = L1.integral (hf.to_L1 f) := -dif_pos hf +@dif_pos _ (id _) hf _ _ _ lemma integral_eq_set_to_fun (f : α → E) : ∫ a, f a ∂μ = set_to_fun μ (weighted_smul μ) (dominated_fin_meas_additive_weighted_smul μ) f := @@ -717,7 +725,7 @@ lemma L1.integral_eq_integral (f : α →₁[μ] E) : L1.integral f = ∫ a, f a (L1.set_to_fun_eq_set_to_L1 (dominated_fin_meas_additive_weighted_smul μ) f).symm lemma integral_undef (h : ¬ integrable f μ) : ∫ a, f a ∂μ = 0 := -dif_neg h +@dif_neg _ (id _) h _ _ _ lemma integral_non_ae_strongly_measurable (h : ¬ ae_strongly_measurable f μ) : ∫ a, f a ∂μ = 0 := integral_undef $ not_and_of_not_left _ h diff --git a/src/measure_theory/probability_mass_function/monad.lean b/src/measure_theory/probability_mass_function/monad.lean index 790c00eef72b7..683b1245c9090 100644 --- a/src/measure_theory/probability_mass_function/monad.lean +++ b/src/measure_theory/probability_mass_function/monad.lean @@ -147,7 +147,7 @@ calc (p.bind f).to_outer_measure s ... = ∑' (a : α), ↑(p a) * ∑' (b : β), if b ∈ s then ↑(f a b) else (0 : ℝ≥0∞) : tsum_congr (λ a, congr_arg (λ x, ↑(p a) * x) $ tsum_congr (λ b, by split_ifs; refl)) ... = ∑' (a : α), ↑(p a) * (f a).to_outer_measure s : - tsum_congr (λ a, by rw [to_outer_measure_apply, set.indicator]) + tsum_congr (λ a, by simp only [to_outer_measure_apply, set.indicator_apply]) /-- The measure of a set under `p.bind f` is the sum over `a : α` of the probability of `a` under `p` times the measure of the set under `f a` -/ diff --git a/src/topology/instances/ennreal.lean b/src/topology/instances/ennreal.lean index 1a0b69d972ba8..017ce827a9183 100644 --- a/src/topology/instances/ennreal.lean +++ b/src/topology/instances/ennreal.lean @@ -841,12 +841,7 @@ lemma tsum_mono_subtype (f : α → ℝ≥0∞) {s t : set α} (h : s ⊆ t) : begin simp only [tsum_subtype], apply ennreal.tsum_le_tsum, - assume x, - split_ifs, - { exact le_rfl }, - { exact (h_2 (h h_1)).elim }, - { exact zero_le _ }, - { exact le_rfl } + exact indicator_le_indicator_of_subset h (λ _, zero_le _), end lemma tsum_union_le (f : α → ℝ≥0∞) (s t : set α) : From e5b48f9bf7f31a18a7d87eeedde572ea2e26c7c6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 2 May 2022 13:45:07 +0000 Subject: [PATCH 011/153] chore(model_theory/basic): golf `countable_empty` (#13836) --- src/model_theory/basic.lean | 9 ++++----- 1 file changed, 4 insertions(+), 5 deletions(-) diff --git a/src/model_theory/basic.lean b/src/model_theory/basic.lean index bc389f9a09b6b..5ea9cac2e6323 100644 --- a/src/model_theory/basic.lean +++ b/src/model_theory/basic.lean @@ -169,12 +169,11 @@ lemma encodable.countable [h : encodable L.symbols] : L.countable := ⟨cardinal.encodable_iff.1 ⟨h⟩⟩ +@[simp] lemma empty_card : language.empty.card = 0 := +by simp [card_eq_card_functions_add_card_relations] + instance countable_empty : language.empty.countable := -⟨begin - rw [card_eq_card_functions_add_card_relations, add_le_omega, lift_le_omega, lift_le_omega, - ← cardinal.encodable_iff, ← cardinal.encodable_iff], - exact ⟨⟨sigma.encodable⟩, ⟨sigma.encodable⟩⟩, -end⟩ +⟨by simp⟩ @[priority 100] instance countable.countable_functions [L.countable] : L.countable_functions := From 9d3db5347e8c637ca8c0247bea649ec9fa500d64 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Mon, 2 May 2022 13:45:08 +0000 Subject: [PATCH 012/153] feat(category_theory/preadditive): End X is a division_ring or field when X is simple (#13849) Consequences of Schur's lemma Co-authored-by: Scott Morrison --- src/category_theory/preadditive/schur.lean | 66 ++++++++++++++++------ 1 file changed, 48 insertions(+), 18 deletions(-) diff --git a/src/category_theory/preadditive/schur.lean b/src/category_theory/preadditive/schur.lean index 931c218419558..22566b9338844 100644 --- a/src/category_theory/preadditive/schur.lean +++ b/src/category_theory/preadditive/schur.lean @@ -19,19 +19,13 @@ Second, we prove Schur's lemma for `𝕜`-linear categories with finite dimensio over an algebraically closed field `𝕜`: the hom space `X ⟶ Y` between simple objects `X` and `Y` is at most one dimensional, and is 1-dimensional iff `X` and `Y` are isomorphic. - -## Future work -It might be nice to provide a `division_ring` instance on `End X` when `X` is simple. -This is an easy consequence of the results here, -but may take some care setting up usable instances. -/ namespace category_theory open category_theory.limits -universes v u -variables {C : Type u} [category.{v} C] +variables {C : Type*} [category C] variables [preadditive C] /-- @@ -49,8 +43,8 @@ end As a corollary of Schur's lemma for preadditive categories, any morphism between simple objects is (exclusively) either an isomorphism or zero. -/ -lemma is_iso_iff_nonzero [has_kernels C] {X Y : C} [simple.{v} X] [simple.{v} Y] (f : X ⟶ Y) : - is_iso.{v} f ↔ f ≠ 0 := +lemma is_iso_iff_nonzero [has_kernels C] {X Y : C} [simple X] [simple Y] (f : X ⟶ Y) : + is_iso f ↔ f ≠ 0 := ⟨λ I, begin introI h, @@ -59,16 +53,34 @@ lemma is_iso_iff_nonzero [has_kernels C] {X Y : C} [simple.{v} X] [simple.{v} Y] end, λ w, is_iso_of_hom_simple w⟩ +/-- +In any preadditive category with kernels, +the endomorphisms of a simple object form a division ring. +-/ +noncomputable +instance [has_kernels C] {X : C} [simple X] : division_ring (End X) := +by classical; exact +{ inv := λ f, if h : f = 0 then 0 else by { haveI := is_iso_of_hom_simple h, exact inv f, }, + exists_pair_ne := ⟨𝟙 X, 0, id_nonzero _⟩, + inv_zero := dif_pos rfl, + mul_inv_cancel := λ f h, begin + haveI := is_iso_of_hom_simple h, + convert is_iso.inv_hom_id f, + exact dif_neg h, + end, + ..(infer_instance : ring (End X)) } + open finite_dimensional -variables (𝕜 : Type*) [field 𝕜] +section +variables (𝕜 : Type*) [division_ring 𝕜] /-- Part of **Schur's lemma** for `𝕜`-linear categories: the hom space between two non-isomorphic simple objects is 0-dimensional. -/ lemma finrank_hom_simple_simple_eq_zero_of_not_iso - [has_kernels C] [linear 𝕜 C] {X Y : C} [simple.{v} X] [simple.{v} Y] + [has_kernels C] [linear 𝕜 C] {X Y : C} [simple X] [simple Y] (h : (X ≅ Y) → false): finrank 𝕜 (X ⟶ Y) = 0 := begin @@ -80,6 +92,9 @@ begin exact finrank_zero_of_subsingleton, end +end + +variables (𝕜 : Type*) [field 𝕜] variables [is_alg_closed 𝕜] [linear 𝕜 C] -- In the proof below we have some difficulty using `I : finite_dimensional 𝕜 (X ⟶ X)` @@ -121,15 +136,30 @@ variables [has_kernels C] **Schur's lemma** for endomorphisms in `𝕜`-linear categories. -/ lemma finrank_endomorphism_simple_eq_one - (X : C) [simple.{v} X] [I : finite_dimensional 𝕜 (X ⟶ X)] : + (X : C) [simple X] [I : finite_dimensional 𝕜 (X ⟶ X)] : finrank 𝕜 (X ⟶ X) = 1 := finrank_endomorphism_eq_one 𝕜 is_iso_iff_nonzero lemma endomorphism_simple_eq_smul_id - {X : C} [simple.{v} X] [I : finite_dimensional 𝕜 (X ⟶ X)] (f : X ⟶ X) : + {X : C} [simple X] [I : finite_dimensional 𝕜 (X ⟶ X)] (f : X ⟶ X) : ∃ c : 𝕜, c • 𝟙 X = f := (finrank_eq_one_iff_of_nonzero' (𝟙 X) (id_nonzero X)).mp (finrank_endomorphism_simple_eq_one 𝕜 X) f +/-- +Endomorphisms of a simple object form a field if they are finite dimensional. +This can't be an instance as `𝕜` would be undetermined. +-/ +noncomputable +def field_End_of_finite_dimensional (X : C) [simple X] [I : finite_dimensional 𝕜 (X ⟶ X)] : + field (End X) := +by classical; exact +{ mul_comm := λ f g, begin + obtain ⟨c, rfl⟩ := endomorphism_simple_eq_smul_id 𝕜 f, + obtain ⟨d, rfl⟩ := endomorphism_simple_eq_smul_id 𝕜 g, + simp [←mul_smul, mul_comm c d], + end, + ..(infer_instance : division_ring (End X)) } + /-- **Schur's lemma** for `𝕜`-linear categories: if hom spaces are finite dimensional, then the hom space between simples is at most 1-dimensional. @@ -137,10 +167,10 @@ if hom spaces are finite dimensional, then the hom space between simples is at m See `finrank_hom_simple_simple_eq_one_iff` and `finrank_hom_simple_simple_eq_zero_iff` below for the refinements when we know whether or not the simples are isomorphic. -/ --- We don't really need `[∀ X Y : C, finite_dimensional 𝕜 (X ⟶ Y)]` here, --- just at least one of `[finite_dimensional 𝕜 (X ⟶ X)]` or `[finite_dimensional 𝕜 (Y ⟶ Y)]`. +-- There is a symmetric argument that uses `[finite_dimensional 𝕜 (Y ⟶ Y)]` instead, +-- but we don't bother proving that here. lemma finrank_hom_simple_simple_le_one - (X Y : C) [∀ X Y : C, finite_dimensional 𝕜 (X ⟶ Y)] [simple.{v} X] [simple.{v} Y] : + (X Y : C) [finite_dimensional 𝕜 (X ⟶ X)] [simple X] [simple Y] : finrank 𝕜 (X ⟶ Y) ≤ 1 := begin cases subsingleton_or_nontrivial (X ⟶ Y) with h, @@ -156,7 +186,7 @@ begin end lemma finrank_hom_simple_simple_eq_one_iff - (X Y : C) [∀ X Y : C, finite_dimensional 𝕜 (X ⟶ Y)] [simple.{v} X] [simple.{v} Y] : + (X Y : C) [finite_dimensional 𝕜 (X ⟶ X)] [finite_dimensional 𝕜 (X ⟶ Y)] [simple X] [simple Y] : finrank 𝕜 (X ⟶ Y) = 1 ↔ nonempty (X ≅ Y) := begin fsplit, @@ -173,7 +203,7 @@ begin end lemma finrank_hom_simple_simple_eq_zero_iff - (X Y : C) [∀ X Y : C, finite_dimensional 𝕜 (X ⟶ Y)] [simple.{v} X] [simple.{v} Y] : + (X Y : C) [finite_dimensional 𝕜 (X ⟶ X)] [finite_dimensional 𝕜 (X ⟶ Y)] [simple X] [simple Y] : finrank 𝕜 (X ⟶ Y) = 0 ↔ is_empty (X ≅ Y) := begin rw [← not_nonempty_iff, ← not_congr (finrank_hom_simple_simple_eq_one_iff 𝕜 X Y)], From dbc033949e8ddb7b221bb9fd451d40db35f9c12f Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Mon, 2 May 2022 13:45:09 +0000 Subject: [PATCH 013/153] feat(category_theory/limits/shapes/types): explicit isos (#13854) Requested on Zulip. https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/Relating.20the.20categorical.20product.20and.20the.20normal.20product Co-authored-by: Scott Morrison --- src/category_theory/limits/shapes/types.lean | 107 ++++++++++++++++++- 1 file changed, 104 insertions(+), 3 deletions(-) diff --git a/src/category_theory/limits/shapes/types.lean b/src/category_theory/limits/shapes/types.lean index a9c6840205d67..95c68a53e2b69 100644 --- a/src/category_theory/limits/shapes/types.lean +++ b/src/category_theory/limits/shapes/types.lean @@ -7,6 +7,7 @@ import category_theory.limits.types import category_theory.limits.shapes.products import category_theory.limits.shapes.binary_products import category_theory.limits.shapes.terminal +import tactic.elementwise /-! # Special shapes for limits in `Type`. @@ -27,8 +28,8 @@ giving the expected definitional implementation: * the pullback of `f : X ⟶ Z` and `g : Y ⟶ Z` is the subtype `{ p : X × Y // f p.1 = g p.2 }` of the product -Because these are not intended for use with the `has_limit` API, -we instead construct terms of `limit_data`. +We first construct terms of `is_limit` and `limit_cone`, and then provide isomorphisms with the +types generated by the `has_limit` API. As an example, when setting up the monoidal category structure on `Type` we use the `types_has_terminal` and `types_has_binary_products` instances. @@ -61,13 +62,21 @@ def terminal_limit_cone : limits.limit_cone (functor.empty (Type u)) := π := by tidy, }, is_limit := by tidy, } +/-- The terminal object in `Type u` is `punit`. -/ +noncomputable def terminal_iso : ⊤_ (Type u) ≅ punit := +limit.iso_limit_cone terminal_limit_cone + /-- The category of types has `pempty` as an initial object. -/ -def initial_limit_cone : limits.colimit_cocone (functor.empty (Type u)) := +def initial_colimit_cocone : limits.colimit_cocone (functor.empty (Type u)) := { cocone := { X := pempty, ι := by tidy, }, is_colimit := by tidy, } +/-- The initial object in `Type u` is `punit`. -/ +noncomputable def initial_iso : ⊥_ (Type u) ≅ pempty := +colimit.iso_colimit_cocone initial_colimit_cocone + open category_theory.limits.walking_pair /-- The product type `X × Y` forms a cone for the binary product of `X` and `Y`. -/ @@ -101,6 +110,26 @@ as the binary product of `X` and `Y`. def binary_product_limit_cone (X Y : Type u) : limits.limit_cone (pair X Y) := ⟨_, binary_product_limit X Y⟩ +/-- The categorical binary product in `Type u` is cartesian product. -/ +noncomputable def binary_product_iso (X Y : Type u) : limits.prod X Y ≅ X × Y := +limit.iso_limit_cone (binary_product_limit_cone X Y) + +@[simp, elementwise] lemma binary_product_iso_hom_comp_fst (X Y : Type u) : + (binary_product_iso X Y).hom ≫ prod.fst = limits.prod.fst := +limit.iso_limit_cone_hom_π (binary_product_limit_cone X Y) walking_pair.left + +@[simp, elementwise] lemma binary_product_iso_hom_comp_snd (X Y : Type u) : + (binary_product_iso X Y).hom ≫ prod.snd = limits.prod.snd := +limit.iso_limit_cone_hom_π (binary_product_limit_cone X Y) walking_pair.right + +@[simp, elementwise] lemma binary_product_iso_inv_comp_fst (X Y : Type u) : + (binary_product_iso X Y).inv ≫ limits.prod.fst = prod.fst := +limit.iso_limit_cone_inv_π (binary_product_limit_cone X Y) walking_pair.left + +@[simp, elementwise] lemma binary_product_iso_inv_comp_snd (X Y : Type u) : + (binary_product_iso X Y).inv ≫ limits.prod.snd = prod.snd := +limit.iso_limit_cone_inv_π (binary_product_limit_cone X Y) walking_pair.right + /-- The functor which sends `X, Y` to the product type `X × Y`. -/ -- We add the option `type_md` to tell `@[simps]` to not treat homomorphisms `X ⟶ Y` in `Type*` as -- a function type @@ -148,6 +177,28 @@ as the binary coproduct of `X` and `Y`. def binary_coproduct_colimit_cocone (X Y : Type u) : limits.colimit_cocone (pair X Y) := ⟨_, binary_coproduct_colimit X Y⟩ +/-- The categorical binary coproduct in `Type u` is the sum `X ⊕ Y`. -/ +noncomputable def binary_coproduct_iso (X Y : Type u) : limits.coprod X Y ≅ X ⊕ Y := +colimit.iso_colimit_cocone (binary_coproduct_colimit_cocone X Y) + +open_locale category_theory.Type + +@[simp, elementwise] lemma binary_coproduct_iso_inl_comp_hom (X Y : Type u) : + limits.coprod.inl ≫ (binary_coproduct_iso X Y).hom = sum.inl := +colimit.iso_colimit_cocone_ι_hom (binary_coproduct_colimit_cocone X Y) walking_pair.left + +@[simp, elementwise] lemma binary_coproduct_iso_inr_comp_hom (X Y : Type u) : + limits.coprod.inr ≫ (binary_coproduct_iso X Y).hom = sum.inr := +colimit.iso_colimit_cocone_ι_hom (binary_coproduct_colimit_cocone X Y) walking_pair.right + +@[simp, elementwise] lemma binary_coproduct_iso_inl_comp_inv (X Y : Type u) : + ↾(sum.inl : X ⟶ X ⊕ Y) ≫ (binary_coproduct_iso X Y).inv = limits.coprod.inl := +colimit.iso_colimit_cocone_ι_inv (binary_coproduct_colimit_cocone X Y) walking_pair.left + +@[simp, elementwise] lemma binary_coproduct_iso_inr_comp_inv (X Y : Type u) : + ↾(sum.inr : Y ⟶ X ⊕ Y) ≫ (binary_coproduct_iso X Y).inv = limits.coprod.inr := +colimit.iso_colimit_cocone_ι_inv (binary_coproduct_colimit_cocone X Y) walking_pair.right + /-- The category of types has `Π j, f j` as the product of a type family `f : J → Type`. -/ @@ -159,6 +210,18 @@ def product_limit_cone {J : Type u} (F : J → Type u) : limits.limit_cone (disc { lift := λ s x j, s.π.app j x, uniq' := λ s m w, funext $ λ x, funext $ λ j, (congr_fun (w j) x : _) } } +/-- The categorical product in `Type u` is the type theoretic product `Π j, F j`. -/ +noncomputable def product_iso {J : Type u} (F : J → Type u) : ∏ F ≅ Π j, F j := +limit.iso_limit_cone (product_limit_cone F) + +@[simp, elementwise] lemma product_iso_hom_comp_eval {J : Type u} (F : J → Type u) (j : J) : + (product_iso F).hom ≫ (λ f, f j) = pi.π F j := +rfl + +@[simp, elementwise] lemma product_iso_inv_comp_π {J : Type u} (F : J → Type u) (j : J) : + (product_iso F).inv ≫ pi.π F j = (λ f, f j) := +limit.iso_limit_cone_inv_π (product_limit_cone F) j + /-- The category of types has `Σ j, f j` as the coproduct of a type family `f : J → Type`. -/ @@ -177,6 +240,18 @@ def coproduct_colimit_cocone {J : Type u} (F : J → Type u) : exact this, end }, } +/-- The categorical coproduct in `Type u` is the type theoretic coproduct `Σ j, F j`. -/ +noncomputable def coproduct_iso {J : Type u} (F : J → Type u) : ∐ F ≅ Σ j, F j := +colimit.iso_colimit_cocone (coproduct_colimit_cocone F) + +@[simp, elementwise] lemma coproduct_iso_ι_comp_hom {J : Type u} (F : J → Type u) (j : J) : + sigma.ι F j ≫ (coproduct_iso F).hom = (λ x : F j, (⟨j, x⟩ : Σ j, F j)) := +colimit.iso_colimit_cocone_ι_hom (coproduct_colimit_cocone F) j + +@[simp, elementwise] lemma coproduct_iso_mk_comp_inv {J : Type u} (F : J → Type u) (j : J) : + ↾(λ x : F j, (⟨j, x⟩ : Σ j, F j)) ≫ (coproduct_iso F).inv = sigma.ι F j := +rfl + section fork variables {X Y Z : Type u} (f : X ⟶ Y) {g h : Y ⟶ Z} (w : f ≫ g = f ≫ h) @@ -227,6 +302,20 @@ def equalizer_limit : limits.limit_cone (parallel_pair g h) := rfl, λ m hm, funext $ λ x, subtype.ext (congr_fun hm x)⟩ } +variables (g h) + +/-- The categorical equalizer in `Type u` is `{x : Y // g x = h x}`. -/ +noncomputable def equalizer_iso : equalizer g h ≅ {x : Y // g x = h x} := +limit.iso_limit_cone equalizer_limit + +@[simp, elementwise] lemma equalizer_iso_hom_comp_subtype : + (equalizer_iso g h).hom ≫ subtype.val = equalizer.ι g h := +rfl + +@[simp, elementwise] lemma equalizer_iso_inv_comp_ι : + (equalizer_iso g h).inv ≫ equalizer.ι g h = subtype.val := +limit.iso_limit_cone_inv_π equalizer_limit walking_parallel_pair.zero + end fork section cofork @@ -271,6 +360,18 @@ begin { exact λ hx, ⟨x, hx, rfl⟩ } end +/-- The categorical coequalizer in `Type u` is the quotient by `f g ~ g x`. -/ +noncomputable def coequalizer_iso : coequalizer f g ≅ _root_.quot (coequalizer_rel f g) := +colimit.iso_colimit_cocone (coequalizer_colimit f g) + +@[simp, elementwise] lemma coequalizer_iso_π_comp_hom : + coequalizer.π f g ≫ (coequalizer_iso f g).hom = quot.mk (coequalizer_rel f g) := +colimit.iso_colimit_cocone_ι_hom (coequalizer_colimit f g) walking_parallel_pair.one + +@[simp, elementwise] lemma coequalizer_iso_quot_comp_inv : + ↾(quot.mk (coequalizer_rel f g)) ≫ (coequalizer_iso f g).inv = coequalizer.π f g := +rfl + end cofork section pullback From ad2e9365113b6966af44c9eb3e40e9b70bd4f15f Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 2 May 2022 13:45:10 +0000 Subject: [PATCH 014/153] feat(topology/homeomorph): add `(co)map_cocompact` (#13861) Also rename `filter.comap_cocompact` to `filter.comap_cocompact_le`. --- src/number_theory/modular.lean | 2 +- src/topology/homeomorph.lean | 8 ++++++++ src/topology/subset_properties.lean | 4 ++-- 3 files changed, 11 insertions(+), 3 deletions(-) diff --git a/src/number_theory/modular.lean b/src/number_theory/modular.lean index b8310760d4f24..7f2d5c9ba7599 100644 --- a/src/number_theory/modular.lean +++ b/src/number_theory/modular.lean @@ -226,7 +226,7 @@ begin { simp only [continuous_pi_iff, fin.forall_fin_two], have : ∀ c : ℝ, continuous (λ x : ℝ, c) := λ c, continuous_const, exact ⟨⟨continuous_id, @this (-1 : ℤ)⟩, ⟨this (cd 0), this (cd 1)⟩⟩ }, - refine filter.tendsto.of_tendsto_comp _ (comap_cocompact hmB), + refine filter.tendsto.of_tendsto_comp _ (comap_cocompact_le hmB), let f₁ : SL(2, ℤ) → matrix (fin 2) (fin 2) ℝ := λ g, matrix.map (↑g : matrix _ _ ℤ) (coe : ℤ → ℝ), have cocompact_ℝ_to_cofinite_ℤ_matrix : diff --git a/src/topology/homeomorph.lean b/src/topology/homeomorph.lean index 3529087fd6330..eaf7605b21f74 100644 --- a/src/topology/homeomorph.lean +++ b/src/topology/homeomorph.lean @@ -174,6 +174,14 @@ h.embedding.is_compact_iff_is_compact_image.symm lemma compact_preimage {s : set β} (h : α ≃ₜ β) : is_compact (h ⁻¹' s) ↔ is_compact s := by rw ← image_symm; exact h.symm.compact_image +@[simp] lemma comap_cocompact (h : α ≃ₜ β) : comap h (cocompact β) = cocompact α := +(comap_cocompact_le h.continuous).antisymm $ + (has_basis_cocompact.le_basis_iff (has_basis_cocompact.comap h)).2 $ λ K hK, + ⟨h ⁻¹' K, h.compact_preimage.2 hK, subset.rfl⟩ + +@[simp] lemma map_cocompact (h : α ≃ₜ β) : map h (cocompact α) = cocompact β := +by rw [← h.comap_cocompact, map_comap_of_surjective h.surjective] + protected lemma compact_space [compact_space α] (h : α ≃ₜ β) : compact_space β := { compact_univ := by { rw [← image_univ_of_surjective h.surjective, h.compact_image], apply compact_space.compact_univ } } diff --git a/src/topology/subset_properties.lean b/src/topology/subset_properties.lean index e5dae21ad05c3..4a642f95dfc9e 100644 --- a/src/topology/subset_properties.lean +++ b/src/topology/subset_properties.lean @@ -759,7 +759,7 @@ fintype_of_univ_finite (hf.finite_of_compact hne) /-- The comap of the cocompact filter on `β` by a continuous function `f : α → β` is less than or equal to the cocompact filter on `α`. This is a reformulation of the fact that images of compact sets are compact. -/ -lemma filter.comap_cocompact {f : α → β} (hf : continuous f) : +lemma filter.comap_cocompact_le {f : α → β} (hf : continuous f) : (filter.cocompact β).comap f ≤ filter.cocompact α := begin rw (filter.has_basis_cocompact.comap f).le_basis_iff filter.has_basis_cocompact, @@ -974,7 +974,7 @@ type `Π d, κ d` the `filter.Coprod` of filters `filter.cocompact` on `κ d`. - lemma filter.Coprod_cocompact {δ : Type*} {κ : δ → Type*} [Π d, topological_space (κ d)] : filter.Coprod (λ d, filter.cocompact (κ d)) = filter.cocompact (Π d, κ d) := begin - refine le_antisymm (supr_le $ λ i, filter.comap_cocompact (continuous_apply i)) _, + refine le_antisymm (supr_le $ λ i, filter.comap_cocompact_le (continuous_apply i)) _, refine compl_surjective.forall.2 (λ s H, _), simp only [compl_mem_Coprod, filter.mem_cocompact, compl_subset_compl, image_subset_iff] at H ⊢, choose K hKc htK using H, From 384a7a39489e613464e05519f2a9324301822f59 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Mon, 2 May 2022 13:45:11 +0000 Subject: [PATCH 015/153] chore(.github/workflows/merge_conflicts.yaml): use separate token (#13884) Co-authored-by: Scott Morrison --- .github/workflows/merge_conflicts.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.github/workflows/merge_conflicts.yml b/.github/workflows/merge_conflicts.yml index 650ee9e46ce84..0fd7c969ce3d0 100644 --- a/.github/workflows/merge_conflicts.yml +++ b/.github/workflows/merge_conflicts.yml @@ -17,4 +17,4 @@ jobs: uses: eps1lon/actions-label-merge-conflict@releases/2.x with: dirtyLabel: "merge-conflict" - repoToken: "${{ secrets.TRIAGE_TOKEN }}" + repoToken: "${{ secrets.MERGE_CONFLICTS_TOKEN }}" From 64bc02c5dec469f8b84c21bc4e49c39ee1b5edcb Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Mon, 2 May 2022 15:58:14 +0000 Subject: [PATCH 016/153] feat(ring_theory/dedekind_domain): Chinese remainder theorem for Dedekind domains (#13067) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The general Chinese remainder theorem states `R / I` is isomorphic to a product of `R / (P i ^ e i)`, where `P i` are comaximal, i.e. `P i + P j = 1` for `i ≠ j`, and the infimum of all `P i` is `I`. In a Dedekind domain the theorem can be stated in a more friendly way, namely that the `P i` are the factors (in the sense of a unique factorization domain) of `I`. This PR provides two ways of doing so, and includes some more lemmas on the ideals in a Dedekind domain. --- src/ring_theory/dedekind_domain/ideal.lean | 168 ++++++++++++++++++++- 1 file changed, 166 insertions(+), 2 deletions(-) diff --git a/src/ring_theory/dedekind_domain/ideal.lean b/src/ring_theory/dedekind_domain/ideal.lean index 63aef98528ce7..464db3d00ec61 100644 --- a/src/ring_theory/dedekind_domain/ideal.lean +++ b/src/ring_theory/dedekind_domain/ideal.lean @@ -3,10 +3,10 @@ Copyright (c) 2020 Kenji Nakagawa. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenji Nakagawa, Anne Baanen, Filippo A. E. Nuccio -/ +import algebra.algebra.subalgebra.pointwise import algebraic_geometry.prime_spectrum.noetherian -import ring_theory.fractional_ideal import ring_theory.dedekind_domain.basic -import algebra.algebra.subalgebra.pointwise +import ring_theory.fractional_ideal /-! # Dedekind domains and ideals @@ -857,3 +857,167 @@ end end is_dedekind_domain end height_one_spectrum + +section chinese_remainder + +open ideal unique_factorization_monoid +open_locale big_operators + +variables {R} + +lemma ring.dimension_le_one.prime_le_prime_iff_eq (h : ring.dimension_le_one R) + {P Q : ideal R} [hP : P.is_prime] [hQ : Q.is_prime] (hP0 : P ≠ ⊥) : + P ≤ Q ↔ P = Q := +⟨(h P hP0 hP).eq_of_le hQ.ne_top, eq.le⟩ + +lemma ideal.coprime_of_no_prime_ge {I J : ideal R} (h : ∀ P, I ≤ P → J ≤ P → ¬ is_prime P) : + I ⊔ J = ⊤ := +begin + by_contra hIJ, + obtain ⟨P, hP, hIJ⟩ := ideal.exists_le_maximal _ hIJ, + exact h P (le_trans le_sup_left hIJ) (le_trans le_sup_right hIJ) hP.is_prime +end + +section dedekind_domain + +variables {R} [is_domain R] [is_dedekind_domain R] + +lemma ideal.is_prime.mul_mem_pow (I : ideal R) [hI : I.is_prime] {a b : R} {n : ℕ} + (h : a * b ∈ I^n) : a ∈ I ∨ b ∈ I^n := +begin + cases n, { simp }, + by_cases hI0 : I = ⊥, { simpa [pow_succ, hI0] using h }, + simp only [← submodule.span_singleton_le_iff_mem, ideal.submodule_span_eq, ← ideal.dvd_iff_le, + ← ideal.span_singleton_mul_span_singleton] at h ⊢, + by_cases ha : I ∣ span {a}, + { exact or.inl ha }, + rw mul_comm at h, + exact or.inr (prime.pow_dvd_of_dvd_mul_right ((ideal.prime_iff_is_prime hI0).mpr hI) _ ha h), +end + +section + +open_locale classical + +lemma ideal.count_normalized_factors_eq {p x : ideal R} (hp0 : p ≠ ⊥) [hp : p.is_prime] {n : ℕ} + (hle : x ≤ p^n) (hlt : ¬ (x ≤ p^(n+1))) : + (normalized_factors x).count p = n := +count_normalized_factors_eq ((ideal.prime_iff_is_prime hp0).mpr hp).irreducible (normalize_eq _) + (ideal.dvd_iff_le.mpr hle) (mt ideal.le_of_dvd hlt) + +end + +lemma ideal.le_mul_of_no_prime_factors + {I J K : ideal R} (coprime : ∀ P, J ≤ P → K ≤ P → ¬ is_prime P) (hJ : I ≤ J) (hK : I ≤ K) : + I ≤ J * K := +begin + simp only [← ideal.dvd_iff_le] at coprime hJ hK ⊢, + by_cases hJ0 : J = 0, + { simpa only [hJ0, zero_mul] using hJ }, + obtain ⟨I', rfl⟩ := hK, + rw mul_comm, + exact mul_dvd_mul_left K + (unique_factorization_monoid.dvd_of_dvd_mul_right_of_no_prime_factors hJ0 + (λ P hPJ hPK, mt ideal.is_prime_of_prime (coprime P hPJ hPK)) + hJ) +end + +lemma ideal.le_of_pow_le_prime {I P : ideal R} [hP : P.is_prime] {n : ℕ} (h : I^n ≤ P) : I ≤ P := +begin + by_cases hP0 : P = ⊥, + { simp only [hP0, le_bot_iff] at ⊢ h, + exact pow_eq_zero h }, + rw ← ideal.dvd_iff_le at ⊢ h, + exact ((ideal.prime_iff_is_prime hP0).mpr hP).dvd_of_dvd_pow h +end + +lemma ideal.pow_le_prime_iff {I P : ideal R} [hP : P.is_prime] {n : ℕ} (hn : n ≠ 0) : + I^n ≤ P ↔ I ≤ P := +⟨ideal.le_of_pow_le_prime, λ h, trans (ideal.pow_le_self hn) h⟩ + +lemma ideal.prod_le_prime {ι : Type*} {s : finset ι} {f : ι → ideal R} {P : ideal R} + [hP : P.is_prime] : + ∏ i in s, f i ≤ P ↔ ∃ i ∈ s, f i ≤ P := +begin + by_cases hP0 : P = ⊥, + { simp only [hP0, le_bot_iff], + rw [← ideal.zero_eq_bot, finset.prod_eq_zero_iff] }, + simp only [← ideal.dvd_iff_le], + exact ((ideal.prime_iff_is_prime hP0).mpr hP).dvd_finset_prod_iff _ +end + +/-- The intersection of distinct prime powers in a Dedekind domain is the product of these +prime powers. -/ +lemma is_dedekind_domain.inf_prime_pow_eq_prod {ι : Type*} + (s : finset ι) (f : ι → ideal R) (e : ι → ℕ) + (prime : ∀ i ∈ s, prime (f i)) (coprime : ∀ i j ∈ s, i ≠ j → f i ≠ f j) : + s.inf (λ i, f i ^ e i) = ∏ i in s, f i ^ e i := +begin + letI := classical.dec_eq ι, + revert prime coprime, + refine s.induction _ _, + { simp }, + intros a s ha ih prime coprime, + specialize ih (λ i hi, prime i (finset.mem_insert_of_mem hi)) + (λ i hi j hj, coprime i (finset.mem_insert_of_mem hi) j (finset.mem_insert_of_mem hj)), + rw [finset.inf_insert, finset.prod_insert ha, ih], + refine le_antisymm (ideal.le_mul_of_no_prime_factors _ inf_le_left inf_le_right) ideal.mul_le_inf, + intros P hPa hPs hPp, + haveI := hPp, + obtain ⟨b, hb, hPb⟩ := ideal.prod_le_prime.mp hPs, + haveI := ideal.is_prime_of_prime (prime a (finset.mem_insert_self a s)), + haveI := ideal.is_prime_of_prime (prime b (finset.mem_insert_of_mem hb)), + refine coprime a (finset.mem_insert_self a s) b (finset.mem_insert_of_mem hb) _ + (((is_dedekind_domain.dimension_le_one.prime_le_prime_iff_eq _).mp + (ideal.le_of_pow_le_prime hPa)).trans + ((is_dedekind_domain.dimension_le_one.prime_le_prime_iff_eq _).mp + (ideal.le_of_pow_le_prime hPb)).symm), + { unfreezingI { rintro rfl }, contradiction }, + { exact (prime a (finset.mem_insert_self a s)).ne_zero }, + { exact (prime b (finset.mem_insert_of_mem hb)).ne_zero }, +end + +/-- **Chinese remainder theorem** for a Dedekind domain: if the ideal `I` factors as +`∏ i, P i ^ e i`, then `R ⧸ I` factors as `Π i, R ⧸ (P i ^ e i)`. -/ +noncomputable def is_dedekind_domain.quotient_equiv_pi_of_prod_eq {ι : Type*} [fintype ι] + (I : ideal R) (P : ι → ideal R) (e : ι → ℕ) + (prime : ∀ i, prime (P i)) (coprime : ∀ i j, i ≠ j → P i ≠ P j) (prod_eq : (∏ i, P i ^ e i) = I) : + R ⧸ I ≃+* Π i, R ⧸ (P i ^ e i) := +(ideal.quot_equiv_of_eq (by { simp only [← prod_eq, finset.inf_eq_infi, finset.mem_univ, cinfi_pos, + ← is_dedekind_domain.inf_prime_pow_eq_prod _ _ _ (λ i _, prime i) (λ i _ j _, coprime i j)] })) + .trans $ +ideal.quotient_inf_ring_equiv_pi_quotient _ (λ i j hij, ideal.coprime_of_no_prime_ge (begin + intros P hPi hPj hPp, + haveI := hPp, + haveI := ideal.is_prime_of_prime (prime i), haveI := ideal.is_prime_of_prime (prime j), + exact coprime i j hij + (((is_dedekind_domain.dimension_le_one.prime_le_prime_iff_eq (prime i).ne_zero).mp + (ideal.le_of_pow_le_prime hPi)).trans + ((is_dedekind_domain.dimension_le_one.prime_le_prime_iff_eq (prime j).ne_zero).mp + (ideal.le_of_pow_le_prime hPj)).symm) +end)) + +open_locale classical + +/-- **Chinese remainder theorem** for a Dedekind domain: `R ⧸ I` factors as `Π i, R ⧸ (P i ^ e i)`, +where `P i` ranges over the prime factors of `I` and `e i` over the multiplicities. -/ +noncomputable def is_dedekind_domain.quotient_equiv_pi_factors {I : ideal R} (hI : I ≠ ⊥) : + R ⧸ I ≃+* Π (P : (factors I).to_finset), R ⧸ ((P : ideal R) ^ (factors I).count P) := +is_dedekind_domain.quotient_equiv_pi_of_prod_eq _ _ _ + (λ (P : (factors I).to_finset), prime_of_factor _ (multiset.mem_to_finset.mp P.prop)) + (λ i j hij, subtype.coe_injective.ne hij) + (calc ∏ (P : (factors I).to_finset), (P : ideal R) ^ (factors I).count (P : ideal R) + = ∏ P in (factors I).to_finset, P ^ (factors I).count P + : (factors I).to_finset.prod_coe_sort (λ P, P ^ (factors I).count P) + ... = ((factors I).map (λ P, P)).prod : (finset.prod_multiset_map_count (factors I) id).symm + ... = (factors I).prod : by rw multiset.map_id' + ... = I : (@associated_iff_eq (ideal R) _ ideal.unique_units _ _).mp (factors_prod hI)) + +@[simp] lemma is_dedekind_domain.quotient_equiv_pi_factors_mk {I : ideal R} (hI : I ≠ ⊥) + (x : R) : is_dedekind_domain.quotient_equiv_pi_factors hI (ideal.quotient.mk I x) = + λ P, ideal.quotient.mk _ x := +rfl + +end dedekind_domain + +end chinese_remainder From 90418df51fdf2c642c160cbe80ab17da15620a72 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Mon, 2 May 2022 15:58:14 +0000 Subject: [PATCH 017/153] feat(linear_algebra/finite_dimensional): `finite_dimensional_iff_of_rank_eq_nsmul` (#13357) If `V` has a dimension that is a scalar multiple of the dimension of `W`, then `V` is finite dimensional iff `W` is. --- src/linear_algebra/finite_dimensional.lean | 7 +++++++ src/set_theory/cardinal/basic.lean | 15 +++++++++++++++ 2 files changed, 22 insertions(+) diff --git a/src/linear_algebra/finite_dimensional.lean b/src/linear_algebra/finite_dimensional.lean index bc69e4e9e9ba5..2ee8800d2efe6 100644 --- a/src/linear_algebra/finite_dimensional.lean +++ b/src/linear_algebra/finite_dimensional.lean @@ -207,6 +207,13 @@ lemma fact_finite_dimensional_of_finrank_eq_succ {K V : Type*} [field K] [add_co finite_dimensional K V := finite_dimensional_of_finrank $ by convert nat.succ_pos n; apply fact.out +lemma finite_dimensional_iff_of_rank_eq_nsmul + {K V W : Type*} [field K] [add_comm_group V] [add_comm_group W] [module K V] [module K W] + {n : ℕ} (hn : n ≠ 0) (hVW : module.rank K V = n • module.rank K W) : + finite_dimensional K V ↔ finite_dimensional K W := +by simp only [finite_dimensional, ← is_noetherian.iff_fg, is_noetherian.iff_dim_lt_omega, hVW, + cardinal.nsmul_lt_omega_iff_of_ne_zero hn] + /-- If a vector space has a finite basis, then its dimension is equal to the cardinality of the basis. -/ lemma finrank_eq_card_basis {ι : Type w} [fintype ι] (h : basis ι K V) : diff --git a/src/set_theory/cardinal/basic.lean b/src/set_theory/cardinal/basic.lean index b805d1ef2fffc..693c3ec06f201 100644 --- a/src/set_theory/cardinal/basic.lean +++ b/src/set_theory/cardinal/basic.lean @@ -988,6 +988,21 @@ lemma add_lt_omega_iff {a b : cardinal} : a + b < ω ↔ a < ω ∧ b < ω := lemma omega_le_add_iff {a b : cardinal} : ω ≤ a + b ↔ ω ≤ a ∨ ω ≤ b := by simp only [← not_lt, add_lt_omega_iff, not_and_distrib] +/-- See also `cardinal.nsmul_lt_omega_iff_of_ne_zero` if you already have `n ≠ 0`. -/ +lemma nsmul_lt_omega_iff {n : ℕ} {a : cardinal} : n • a < ω ↔ n = 0 ∨ a < ω := +begin + cases n, + { simpa using nat_lt_omega 0 }, + simp only [nat.succ_ne_zero, false_or], + induction n with n ih, + { simp }, + rw [succ_nsmul, add_lt_omega_iff, ih, and_self] +end + +/-- See also `cardinal.nsmul_lt_omega_iff` for a hypothesis-free version. -/ +lemma nsmul_lt_omega_iff_of_ne_zero {n : ℕ} {a : cardinal} (h : n ≠ 0) : n • a < ω ↔ a < ω := +nsmul_lt_omega_iff.trans $ or_iff_right h + theorem mul_lt_omega {a b : cardinal} (ha : a < ω) (hb : b < ω) : a * b < ω := match a, b, lt_omega.1 ha, lt_omega.1 hb with | _, _, ⟨m, rfl⟩, ⟨n, rfl⟩ := by rw [← nat.cast_mul]; apply nat_lt_omega From 65843cdc05700c540d337d4ef3b9d20e44841e73 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 2 May 2022 15:58:16 +0000 Subject: [PATCH 018/153] feat(analysis/matrix): provide a normed_algebra structure on matrices (#13518) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This is one of the final pieces needed to defining the matrix exponential. It would be nice to show: ```lean lemma l1_linf_norm_to_matrix [nondiscrete_normed_field R] [decidable_eq n] (f : (n → R) →L[R] (m → R)) : ∥linear_map.to_matrix' (↑f : (n → R) →ₗ[R] (m → R))∥ = ∥f∥ := ``` but its not clear to me under what generality it holds. --- src/analysis/matrix.lean | 216 ++++++++++++++++++++++++++++++++++++--- 1 file changed, 201 insertions(+), 15 deletions(-) diff --git a/src/analysis/matrix.lean b/src/analysis/matrix.lean index 34ec40dadf1e9..cd12ceb35994d 100644 --- a/src/analysis/matrix.lean +++ b/src/analysis/matrix.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2021 Heather Macbeth. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Heather Macbeth +Authors: Heather Macbeth, Eric Wieser -/ import analysis.normed_space.basic import analysis.normed_space.pi_Lp @@ -10,32 +10,49 @@ import analysis.inner_product_space.pi_L2 /-! # Matrices as a normed space -In this file we provide the following non-instances on matrices, using the elementwise norm: +In this file we provide the following non-instances for norms on matrices: -* `matrix.semi_normed_group` -* `matrix.normed_group` -* `matrix.normed_space` +* The elementwise norm: -These are not declared as instances because there are several natural choices for defining the norm -of a matrix. + * `matrix.semi_normed_group` + * `matrix.normed_group` + * `matrix.normed_space` + +* The Frobenius norm: + + * `matrix.frobenius_semi_normed_group` + * `matrix.frobenius_normed_group` + * `matrix.frobenius_normed_space` + * `matrix.frobenius_normed_ring` + * `matrix.frobenius_normed_algebra` + +* The $L^\infty$ operator norm: -We also provide definitions of the frobenius norm, again not declared as instances: + * `matrix.linfty_op_semi_normed_group` + * `matrix.linfty_op_normed_group` + * `matrix.linfty_op_normed_space` + * `matrix.linfty_op_non_unital_semi_normed_ring` + * `matrix.linfty_op_semi_normed_ring` + * `matrix.linfty_op_non_unital_normed_ring` + * `matrix.linfty_op_normed_ring` + * `matrix.linfty_op_normed_algebra` -* `matrix.frobenius_semi_normed_group` -* `matrix.frobenius_normed_group` -* `matrix.frobenius_normed_space` -* `matrix.frobenius_normed_ring` -* `matrix.frobenius_normed_algebra` +These are not declared as instances because there are several natural choices for defining the norm +of a matrix. -/ noncomputable theory -open_locale nnreal matrix +open_locale big_operators nnreal matrix namespace matrix variables {R l m n α β : Type*} [fintype l] [fintype m] [fintype n] +/-! ### The elementwise supremum norm -/ + +section linf_linf + section semi_normed_group variables [semi_normed_group α] [semi_normed_group β] @@ -130,7 +147,175 @@ pi.normed_space end normed_space -/-! ### The Frobenius norm on matrices +end linf_linf + +/-! ### The $L_\infty$ operator norm + +This section defines the matrix norm $\|A\|_\infty = \operatorname{sup}_i (\sum_j \|A_{ij}\|)$. + +Note that this is equivalent to the operator norm, considering $A$ as a linear map between two +$L^\infty$ spaces. +-/ +section linfty_op + +/-- Seminormed group instance (using sup norm of L1 norm) for matrices over a seminormed group. Not +declared as an instance because there are several natural choices for defining the norm of a +matrix. -/ +local attribute [instance] +protected def linfty_op_semi_normed_group [semi_normed_group α] : + semi_normed_group (matrix m n α) := +(by apply_instance : semi_normed_group (m → pi_Lp 1 (λ j : n, α))) + +/-- Normed group instance (using sup norm of L1 norm) for matrices over a normed ring. Not +declared as an instance because there are several natural choices for defining the norm of a +matrix. -/ +local attribute [instance] +protected def linfty_op_normed_group [normed_group α] : + normed_group (matrix m n α) := +(by apply_instance : normed_group (m → pi_Lp 1 (λ j : n, α))) + +/-- Normed space instance (using sup norm of L1 norm) for matrices over a normed space. Not +declared as an instance because there are several natural choices for defining the norm of a +matrix. -/ +local attribute [instance] +protected def linfty_op_normed_space [normed_field R] [semi_normed_group α] [normed_space R α] : + normed_space R (matrix m n α) := +(by apply_instance : normed_space R (m → pi_Lp 1 (λ j : n, α))) + +section semi_normed_group +variables [semi_normed_group α] + +lemma linfty_op_norm_def (A : matrix m n α) : + ∥A∥ = ((finset.univ : finset m).sup (λ i : m, ∑ j : n, ∥A i j∥₊) : ℝ≥0) := +by simp_rw [pi.norm_def, pi_Lp.nnnorm_eq, div_one, nnreal.rpow_one] + +lemma linfty_op_nnnorm_def (A : matrix m n α) : + ∥A∥₊ = (finset.univ : finset m).sup (λ i : m, ∑ j : n, ∥A i j∥₊) := +subtype.ext $ linfty_op_norm_def A + +@[simp] lemma linfty_op_nnnorm_col (v : m → α) : + ∥col v∥₊ = ∥v∥₊ := +begin + rw [linfty_op_nnnorm_def, pi.nnnorm_def], + simp, +end + +@[simp] lemma linfty_op_norm_col (v : m → α) : + ∥col v∥ = ∥v∥ := +congr_arg coe $ linfty_op_nnnorm_col v + +@[simp] lemma linfty_op_nnnorm_row (v : n → α) : + ∥row v∥₊ = ∑ i, ∥v i∥₊ := +by simp [linfty_op_nnnorm_def] + +@[simp] lemma linfty_op_norm_row (v : n → α) : + ∥row v∥ = ∑ i, ∥v i∥ := +(congr_arg coe $ linfty_op_nnnorm_row v).trans $ by simp [nnreal.coe_sum] + +@[simp] +lemma linfty_op_nnnorm_diagonal [decidable_eq m] (v : m → α) : + ∥diagonal v∥₊ = ∥v∥₊ := +begin + rw [linfty_op_nnnorm_def, pi.nnnorm_def], + congr' 1 with i : 1, + refine (finset.sum_eq_single_of_mem _ (finset.mem_univ i) $ λ j hj hij, _).trans _, + { rw [diagonal_apply_ne' _ hij, nnnorm_zero] }, + { rw [diagonal_apply_eq] }, +end + +@[simp] +lemma linfty_op_norm_diagonal [decidable_eq m] (v : m → α) : + ∥diagonal v∥ = ∥v∥ := +congr_arg coe $ linfty_op_nnnorm_diagonal v + +end semi_normed_group + +section non_unital_semi_normed_ring +variables [non_unital_semi_normed_ring α] + +lemma linfty_op_nnnorm_mul (A : matrix l m α) (B : matrix m n α) : ∥A ⬝ B∥₊ ≤ ∥A∥₊ * ∥B∥₊ := +begin + simp_rw [linfty_op_nnnorm_def, matrix.mul_apply], + calc finset.univ.sup (λ i, ∑ k, ∥∑ j, A i j * B j k∥₊) + ≤ finset.univ.sup (λ i, ∑ k j, ∥A i j∥₊ * ∥B j k∥₊) : + finset.sup_mono_fun $ λ i hi, finset.sum_le_sum $ λ k hk, nnnorm_sum_le_of_le _ $ λ j hj, + nnnorm_mul_le _ _ + ... = finset.univ.sup (λ i, ∑ j, (∥A i j∥₊ * ∑ k, ∥B j k∥₊)) : + by simp_rw [@finset.sum_comm _ m n, finset.mul_sum] + ... ≤ finset.univ.sup (λ i, ∑ j, ∥A i j∥₊ * finset.univ.sup (λ i, ∑ j, ∥B i j∥₊)) : + finset.sup_mono_fun $ λ i hi, finset.sum_le_sum $ λ j hj, + mul_le_mul_of_nonneg_left (finset.le_sup hj) (zero_le _) + ... ≤ finset.univ.sup (λ i, ∑ j, ∥A i j∥₊) * finset.univ.sup (λ i, ∑ j, ∥B i j∥₊) : + by simp_rw [←finset.sum_mul, ←nnreal.finset_sup_mul], +end + +lemma linfty_op_norm_mul (A : matrix l m α) (B : matrix m n α) : ∥A ⬝ B∥ ≤ ∥A∥ * ∥B∥ := +linfty_op_nnnorm_mul _ _ + +lemma linfty_op_nnnorm_mul_vec (A : matrix l m α) (v : m → α) : ∥A.mul_vec v∥₊ ≤ ∥A∥₊ * ∥v∥₊ := +begin + rw [←linfty_op_nnnorm_col (A.mul_vec v), ←linfty_op_nnnorm_col v], + exact linfty_op_nnnorm_mul A (col v), +end + +lemma linfty_op_norm_mul_vec (A : matrix l m α) (v : m → α) : ∥matrix.mul_vec A v∥ ≤ ∥A∥ * ∥v∥ := +linfty_op_nnnorm_mul_vec _ _ + +end non_unital_semi_normed_ring + +/-- Seminormed non-unital ring instance (using sup norm of L1 norm) for matrices over a semi normed +non-unital ring. Not declared as an instance because there are several natural choices for defining +the norm of a matrix. -/ +local attribute [instance] +protected def linfty_op_non_unital_semi_normed_ring [non_unital_semi_normed_ring α] : + non_unital_semi_normed_ring (matrix n n α) := +{ norm_mul := linfty_op_norm_mul, + .. matrix.linfty_op_semi_normed_group, + .. matrix.non_unital_ring } + +/-- The `L₁-L∞` norm preserves one on non-empty matrices. Note this is safe as an instance, as it +carries no data. -/ +instance linfty_op_norm_one_class [semi_normed_ring α] [norm_one_class α] [decidable_eq n] + [nonempty n] : norm_one_class (matrix n n α) := +{ norm_one := (linfty_op_norm_diagonal _).trans norm_one } + +/-- Seminormed ring instance (using sup norm of L1 norm) for matrices over a semi normed ring. Not +declared as an instance because there are several natural choices for defining the norm of a +matrix. -/ +local attribute [instance] +protected def linfty_op_semi_normed_ring [semi_normed_ring α] [decidable_eq n] : + semi_normed_ring (matrix n n α) := +{ .. matrix.linfty_op_non_unital_semi_normed_ring, + .. matrix.ring } + +/-- Normed non-unital ring instance (using sup norm of L1 norm) for matrices over a normed +non-unital ring. Not declared as an instance because there are several natural choices for defining +the norm of a matrix. -/ +local attribute [instance] +protected def linfty_op_non_unital_normed_ring [non_unital_normed_ring α] : + non_unital_normed_ring (matrix n n α) := +{ ..matrix.linfty_op_non_unital_semi_normed_ring } + +/-- Normed ring instance (using sup norm of L1 norm) for matrices over a normed ring. Not +declared as an instance because there are several natural choices for defining the norm of a +matrix. -/ +local attribute [instance] +protected def linfty_op_normed_ring [normed_ring α] [decidable_eq n] : + normed_ring (matrix n n α) := +{ ..matrix.linfty_op_semi_normed_ring } + +/-- Normed algebra instance (using sup norm of L1 norm) for matrices over a normed algebra. Not +declared as an instance because there are several natural choices for defining the norm of a +matrix. -/ +local attribute [instance] +protected def linfty_op_normed_algebra [normed_field R] [semi_normed_ring α] [normed_algebra R α] + [decidable_eq n] : + normed_algebra R (matrix n n α) := +{ ..matrix.linfty_op_normed_space } + +end linfty_op + +/-! ### The Frobenius norm This is defined as $\|A\| = \sqrt{\sum_ij \|A_ij\|^2}$. When the matrix is over the real or complex numbers, this norm is submultiplicative. @@ -267,4 +452,5 @@ def frobenius_normed_algebra [decidable_eq m] [normed_field R] [normed_algebra R end is_R_or_C end frobenius + end matrix From af11e154f8917054fd6e73f21a82893874b3d19e Mon Sep 17 00:00:00 2001 From: Oliver Nash Date: Mon, 2 May 2022 15:58:17 +0000 Subject: [PATCH 019/153] feat(algebra/big_operators/finprod): add lemma `finprod_eq_prod_of_mul_support_to_finset_subset'` (#13801) Formalized as part of the Sphere Eversion project. --- src/algebra/big_operators/finprod.lean | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/src/algebra/big_operators/finprod.lean b/src/algebra/big_operators/finprod.lean index ebb13bc45c760..4e2771dd7af7e 100644 --- a/src/algebra/big_operators/finprod.lean +++ b/src/algebra/big_operators/finprod.lean @@ -285,6 +285,15 @@ end ∏ᶠ i, f i = ∏ i in s, f i := finprod_eq_prod_of_mul_support_subset _ $ λ x hx, h $ hf.mem_to_finset.2 hx +@[to_additive] lemma finprod_eq_finset_prod_of_mul_support_subset + (f : α → M) {s : finset α} (h : mul_support f ⊆ (s : set α)) : + ∏ᶠ i, f i = ∏ i in s, f i := +begin + have h' : (s.finite_to_set.subset h).to_finset ⊆ s, + { simpa [← finset.coe_subset, set.coe_to_finset], }, + exact finprod_eq_prod_of_mul_support_to_finset_subset _ _ h', +end + @[to_additive] lemma finprod_def (f : α → M) [decidable (mul_support f).finite] : ∏ᶠ i : α, f i = if h : (mul_support f).finite then ∏ i in h.to_finset, f i else 1 := begin From 917b5275f33ff60aa0e1f60bf3bb6ef8203a50a5 Mon Sep 17 00:00:00 2001 From: kkytola Date: Mon, 2 May 2022 17:28:45 +0000 Subject: [PATCH 020/153] feat(topology/metric_space/thickened_indicator): Add definition and lemmas about thickened indicators. (#13481) Add thickened indicators, to be used for the proof of the portmanteau theorem. --- src/data/real/ennreal.lean | 3 + src/topology/instances/ennreal.lean | 8 + .../metric_space/thickened_indicator.lean | 235 ++++++++++++++++++ 3 files changed, 246 insertions(+) create mode 100644 src/topology/metric_space/thickened_indicator.lean diff --git a/src/data/real/ennreal.lean b/src/data/real/ennreal.lean index 446b17bb3b5c3..001e3031e1c04 100644 --- a/src/data/real/ennreal.lean +++ b/src/data/real/ennreal.lean @@ -1180,6 +1180,9 @@ begin exact le_div_iff_mul_le (or.inl (mt coe_eq_coe.1 h)) (or.inl coe_ne_top) end +lemma div_le_div {a b c d : ℝ≥0∞} (hab : a ≤ b) (hdc : d ≤ c) : a / c ≤ b / d := +div_eq_mul_inv b d ▸ div_eq_mul_inv a c ▸ ennreal.mul_le_mul hab (ennreal.inv_le_inv.mpr hdc) + lemma mul_inv_cancel (h0 : a ≠ 0) (ht : a ≠ ∞) : a * a⁻¹ = 1 := begin lift a to ℝ≥0 using ht, diff --git a/src/topology/instances/ennreal.lean b/src/topology/instances/ennreal.lean index 017ce827a9183..389186e1394b0 100644 --- a/src/topology/instances/ennreal.lean +++ b/src/topology/instances/ennreal.lean @@ -371,6 +371,14 @@ continuous_iff_continuous_at.2 $ λ x, ennreal.continuous_at_const_mul (or.inl h protected lemma continuous_mul_const {a : ℝ≥0∞} (ha : a ≠ ⊤) : continuous (λ x, x * a) := continuous_iff_continuous_at.2 $ λ x, ennreal.continuous_at_mul_const (or.inl ha) +protected lemma continuous_div_const (c : ℝ≥0∞) (c_ne_zero : c ≠ 0) : + continuous (λ (x : ℝ≥0∞), x / c) := +begin + simp_rw [div_eq_mul_inv, continuous_iff_continuous_at], + intro x, + exact ennreal.continuous_at_mul_const (or.intro_left _ (inv_ne_top.mpr c_ne_zero)), +end + @[continuity] lemma continuous_pow (n : ℕ) : continuous (λ a : ℝ≥0∞, a ^ n) := begin diff --git a/src/topology/metric_space/thickened_indicator.lean b/src/topology/metric_space/thickened_indicator.lean new file mode 100644 index 0000000000000..bc45cf672a098 --- /dev/null +++ b/src/topology/metric_space/thickened_indicator.lean @@ -0,0 +1,235 @@ +/- +Copyright (c) 2022 Kalle Kytölä. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kalle Kytölä +-/ +import data.real.ennreal +import topology.continuous_function.bounded + +/-! +# Thickened indicators + +This file is about thickened indicators of sets in (pseudo e)metric spaces. For a decreasing +sequence of thickening radii tending to 0, the thickened indicators of a closed set form a +decreasing pointwise converging approximation of the indicator function of the set, where the +members of the approximating sequence are nonnegative bounded continuous functions. + +## Main definitions + + * `thickened_indicator_aux δ E`: The `δ`-thickened indicator of a set `E` as an + unbundled `ℝ≥0∞`-valued function. + * `thickened_indicator δ E`: The `δ`-thickened indicator of a set `E` as a bundled + bounded continuous `ℝ≥0`-valued function. + +## Main results + + * For a sequence of thickening radii tending to 0, the `δ`-thickened indicators of a set `E` tend + pointwise to the indicator of `closure E`. + - `thickened_indicator_aux_tendsto_indicator_closure`: The version is for the + unbundled `ℝ≥0∞`-valued functions. + - `thickened_indicator_tendsto_indicator_closure`: The version is for the bundled `ℝ≥0`-valued + bounded continuous functions. + +-/ +noncomputable theory +open_locale classical nnreal ennreal topological_space bounded_continuous_function + +open nnreal ennreal set metric emetric filter + +section thickened_indicator + +variables {α : Type*} [pseudo_emetric_space α] + +/-- The `δ`-thickened indicator of a set `E` is the function that equals `1` on `E` +and `0` outside a `δ`-thickening of `E` and interpolates (continuously) between +these values using `inf_edist _ E`. + +`thickened_indicator_aux` is the unbundled `ℝ≥0∞`-valued function. See `thickened_indicator` +for the (bundled) bounded continuous function with `ℝ≥0`-values. -/ +def thickened_indicator_aux (δ : ℝ) (E : set α) : α → ℝ≥0∞ := +λ (x : α), (1 : ℝ≥0∞) - (inf_edist x E) / (ennreal.of_real δ) + +lemma continuous_thickened_indicator_aux {δ : ℝ} (δ_pos : 0 < δ) (E : set α) : + continuous (thickened_indicator_aux δ E) := +begin + unfold thickened_indicator_aux, + let f := λ (x : α), (⟨1, (inf_edist x E) / (ennreal.of_real δ)⟩ : ℝ≥0 × ℝ≥0∞), + let sub := λ (p : ℝ≥0 × ℝ≥0∞), ((p.1 : ℝ≥0∞) - p.2), + rw (show (λ (x : α), ((1 : ℝ≥0∞)) - (inf_edist x E) / (ennreal.of_real δ)) = sub ∘ f, by refl), + apply (@ennreal.continuous_nnreal_sub 1).comp, + apply (ennreal.continuous_div_const (ennreal.of_real δ) _).comp continuous_inf_edist, + norm_num [δ_pos], +end + +lemma thickened_indicator_aux_le_one (δ : ℝ) (E : set α) (x : α) : + thickened_indicator_aux δ E x ≤ 1 := +by apply @tsub_le_self _ _ _ _ (1 : ℝ≥0∞) + +lemma thickened_indicator_aux_lt_top {δ : ℝ} {E : set α} {x : α} : + thickened_indicator_aux δ E x < ∞ := +lt_of_le_of_lt (thickened_indicator_aux_le_one _ _ _) one_lt_top + +lemma thickened_indicator_aux_closure_eq (δ : ℝ) (E : set α) : + thickened_indicator_aux δ (closure E) = thickened_indicator_aux δ E := +by simp_rw [thickened_indicator_aux, inf_edist_closure] + +lemma thickened_indicator_aux_one (δ : ℝ) (E : set α) {x : α} (x_in_E : x ∈ E) : + thickened_indicator_aux δ E x = 1 := +by simp [thickened_indicator_aux, inf_edist_zero_of_mem x_in_E, tsub_zero] + +lemma thickened_indicator_aux_one_of_mem_closure + (δ : ℝ) (E : set α) {x : α} (x_mem : x ∈ closure E) : + thickened_indicator_aux δ E x = 1 := +by rw [←thickened_indicator_aux_closure_eq, thickened_indicator_aux_one δ (closure E) x_mem] + +lemma thickened_indicator_aux_zero + {δ : ℝ} (δ_pos : 0 < δ) (E : set α) {x : α} (x_out : x ∉ thickening δ E) : + thickened_indicator_aux δ E x = 0 := +begin + rw [thickening, mem_set_of_eq, not_lt] at x_out, + unfold thickened_indicator_aux, + apply le_antisymm _ bot_le, + have key := tsub_le_tsub (@rfl _ (1 : ℝ≥0∞)).le (ennreal.div_le_div x_out rfl.le), + rw [ennreal.div_self (ne_of_gt (ennreal.of_real_pos.mpr δ_pos)) of_real_ne_top] at key, + simpa using key, +end + +lemma thickened_indicator_aux_mono {δ₁ δ₂ : ℝ} (hle : δ₁ ≤ δ₂) (E : set α) : + thickened_indicator_aux δ₁ E ≤ thickened_indicator_aux δ₂ E := +λ _, tsub_le_tsub (@rfl ℝ≥0∞ 1).le (ennreal.div_le_div rfl.le (of_real_le_of_real hle)) + +lemma thickened_indicator_aux_subset (δ : ℝ) {E₁ E₂ : set α} (subset : E₁ ⊆ E₂) : + thickened_indicator_aux δ E₁ ≤ thickened_indicator_aux δ E₂ := +λ _, tsub_le_tsub (@rfl ℝ≥0∞ 1).le (ennreal.div_le_div (inf_edist_anti subset) rfl.le) + +/-- As the thickening radius δ tends to 0, the δ-thickened indicator of a set E (in α) tends +pointwise (i.e., w.r.t. the product topology on `α → ℝ≥0∞`) to the indicator function of the +closure of E. + +This statement is for the unbundled `ℝ≥0∞`-valued functions `thickened_indicator_aux δ E`, see +`thickened_indicator_tendsto_indicator_closure` for the version for bundled `ℝ≥0`-valued +bounded continuous functions. -/ +lemma thickened_indicator_aux_tendsto_indicator_closure + {δseq : ℕ → ℝ} (δseq_lim : tendsto δseq at_top (𝓝 0)) (E : set α) : + tendsto (λ n, (thickened_indicator_aux (δseq n) E)) at_top + (𝓝 (indicator (closure E) (λ x, (1 : ℝ≥0∞)))) := +begin + rw tendsto_pi_nhds, + intro x, + by_cases x_mem_closure : x ∈ closure E, + { simp_rw [thickened_indicator_aux_one_of_mem_closure _ E x_mem_closure], + rw (show (indicator (closure E) (λ _, (1 : ℝ≥0∞))) x = 1, + by simp only [x_mem_closure, indicator_of_mem]), + exact tendsto_const_nhds, }, + { rw (show (closure E).indicator (λ _, (1 : ℝ≥0∞)) x = 0, + by simp only [x_mem_closure, indicator_of_not_mem, not_false_iff]), + rw mem_closure_iff_inf_edist_zero at x_mem_closure, + obtain ⟨ε, ⟨ε_pos, ε_le⟩⟩ : ∃ (ε : ℝ), 0 < ε ∧ ennreal.of_real ε ≤ inf_edist x E, + { by_cases dist_infty : inf_edist x E = ∞, + { rw dist_infty, + use [1, zero_lt_one, le_top], }, + { use (inf_edist x E).to_real, + exact ⟨(to_real_lt_to_real zero_ne_top dist_infty).mpr (pos_iff_ne_zero.mpr x_mem_closure), + of_real_to_real_le⟩, }, }, + rw metric.tendsto_nhds at δseq_lim, + specialize δseq_lim ε ε_pos, + simp only [dist_zero_right, real.norm_eq_abs, eventually_at_top, ge_iff_le] at δseq_lim, + rcases δseq_lim with ⟨N, hN⟩, + apply @tendsto_at_top_of_eventually_const _ _ _ _ _ _ _ N, + intros n n_large, + have key : x ∉ thickening ε E, by rwa [thickening, mem_set_of_eq, not_lt], + refine le_antisymm _ bot_le, + apply (thickened_indicator_aux_mono (lt_of_abs_lt (hN n n_large)).le E x).trans, + exact (thickened_indicator_aux_zero ε_pos E key).le, }, +end + +/-- The `δ`-thickened indicator of a set `E` is the function that equals `1` on `E` +and `0` outside a `δ`-thickening of `E` and interpolates (continuously) between +these values using `inf_edist _ E`. + +`thickened_indicator` is the (bundled) bounded continuous function with `ℝ≥0`-values. +See `thickened_indicator_aux` for the unbundled `ℝ≥0∞`-valued function. -/ +@[simps] def thickened_indicator {δ : ℝ} (δ_pos : 0 < δ) (E : set α) : α →ᵇ ℝ≥0 := +{ to_fun := λ (x : α), (thickened_indicator_aux δ E x).to_nnreal, + continuous_to_fun := begin + apply continuous_on.comp_continuous + continuous_on_to_nnreal (continuous_thickened_indicator_aux δ_pos E), + intro x, + exact (lt_of_le_of_lt (@thickened_indicator_aux_le_one _ _ δ E x) one_lt_top).ne, + end, + map_bounded' := begin + use 2, + intros x y, + rw [nnreal.dist_eq], + apply (abs_sub _ _).trans, + rw [nnreal.abs_eq, nnreal.abs_eq, ←one_add_one_eq_two], + have key := @thickened_indicator_aux_le_one _ _ δ E, + apply add_le_add; + { norm_cast, + refine (to_nnreal_le_to_nnreal ((lt_of_le_of_lt (key _) one_lt_top).ne) one_ne_top).mpr + (key _), }, + end, } + +lemma thickened_indicator.coe_fn_eq_comp {δ : ℝ} (δ_pos : 0 < δ) (E : set α) : + ⇑(thickened_indicator δ_pos E) = ennreal.to_nnreal ∘ thickened_indicator_aux δ E := rfl + +lemma thickened_indicator_le_one {δ : ℝ} (δ_pos : 0 < δ) (E : set α) (x : α) : + thickened_indicator δ_pos E x ≤ 1 := +begin + rw [thickened_indicator.coe_fn_eq_comp], + simpa using (to_nnreal_le_to_nnreal thickened_indicator_aux_lt_top.ne one_ne_top).mpr + (thickened_indicator_aux_le_one δ E x), +end + +lemma thickened_indicator_one_of_mem_closure + {δ : ℝ} (δ_pos : 0 < δ) (E : set α) {x : α} (x_mem : x ∈ closure E) : + thickened_indicator δ_pos E x = 1 := +by rw [thickened_indicator_apply, + thickened_indicator_aux_one_of_mem_closure δ E x_mem, one_to_nnreal] + +lemma thickened_indicator_one {δ : ℝ} (δ_pos : 0 < δ) (E : set α) {x : α} (x_in_E : x ∈ E) : + thickened_indicator δ_pos E x = 1 := +thickened_indicator_one_of_mem_closure _ _ (subset_closure x_in_E) + +lemma thickened_indicator_zero + {δ : ℝ} (δ_pos : 0 < δ) (E : set α) {x : α} (x_out : x ∉ thickening δ E) : + thickened_indicator δ_pos E x = 0 := +by rw [thickened_indicator_apply, thickened_indicator_aux_zero δ_pos E x_out, zero_to_nnreal] + +lemma thickened_indicator_mono {δ₁ δ₂ : ℝ} + (δ₁_pos : 0 < δ₁) (δ₂_pos : 0 < δ₂) (hle : δ₁ ≤ δ₂) (E : set α) : + ⇑(thickened_indicator δ₁_pos E) ≤ thickened_indicator δ₂_pos E := +begin + intro x, + apply (to_nnreal_le_to_nnreal thickened_indicator_aux_lt_top.ne + thickened_indicator_aux_lt_top.ne).mpr, + apply thickened_indicator_aux_mono hle, +end + +lemma thickened_indicator_subset {δ : ℝ} (δ_pos : 0 < δ) {E₁ E₂ : set α} (subset : E₁ ⊆ E₂) : + ⇑(thickened_indicator δ_pos E₁) ≤ thickened_indicator δ_pos E₂ := +λ x, (to_nnreal_le_to_nnreal thickened_indicator_aux_lt_top.ne + thickened_indicator_aux_lt_top.ne).mpr (thickened_indicator_aux_subset δ subset x) + +/-- As the thickening radius δ tends to 0, the δ-thickened indicator of a set E (in α) tends +pointwise to the indicator function of the closure of E. + +Note: This version is for the bundled bounded continuous functions, but the topology is not +the topology on `α →ᵇ ℝ≥0`. Coercions to functions `α → ℝ≥0` are done first, so the topology +instance is the product topology (the topology of pointwise convergence). -/ +lemma thickened_indicator_tendsto_indicator_closure + {δseq : ℕ → ℝ} (δseq_pos : ∀ n, 0 < δseq n) (δseq_lim : tendsto δseq at_top (𝓝 0)) (E : set α) : + tendsto (λ (n : ℕ), (coe_fn : (α →ᵇ ℝ≥0) → (α → ℝ≥0)) (thickened_indicator (δseq_pos n) E)) + at_top (𝓝 (indicator (closure E) (λ x, (1 : ℝ≥0)))) := +begin + have key := thickened_indicator_aux_tendsto_indicator_closure δseq_lim E, + rw tendsto_pi_nhds at *, + intro x, + rw (show indicator (closure E) (λ x, (1 : ℝ≥0)) x + = (indicator (closure E) (λ x, (1 : ℝ≥0∞)) x).to_nnreal, + by refine (congr_fun (comp_indicator_const 1 ennreal.to_nnreal zero_to_nnreal) x).symm), + refine tendsto.comp (tendsto_to_nnreal _) (key x), + by_cases x_mem : x ∈ closure E; simp [x_mem], +end + +end thickened_indicator -- section From 206a5f773bbbaf2adcf320825106097f0f193019 Mon Sep 17 00:00:00 2001 From: kkytola Date: Mon, 2 May 2022 17:28:46 +0000 Subject: [PATCH 021/153] feat(measure_theory/integral/bochner): Add a rewrite lemma saying the ennreal coercion of an integral of a nonnegative function equals the lintegral of the ennreal coercion of the function. (#13701) This PR adds a rewrite lemma `of_real_integral_eq_lintegral_of_real` that is very similar to `lintegral_coe_eq_integral`, but for nonnegative real-valued functions instead of nnreal-valued functions. Co-authored-by: kkytola <39528102+kkytola@users.noreply.github.com> --- src/measure_theory/integral/bochner.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/src/measure_theory/integral/bochner.lean b/src/measure_theory/integral/bochner.lean index 1347b71d05914..2d76699bfd7f1 100644 --- a/src/measure_theory/integral/bochner.lean +++ b/src/measure_theory/integral/bochner.lean @@ -1013,6 +1013,18 @@ begin rw [← lt_top_iff_ne_top], convert hfi.has_finite_integral, ext1 x, rw [nnreal.nnnorm_eq] end +lemma of_real_integral_eq_lintegral_of_real {f : α → ℝ} (hfi : integrable f μ) (f_nn : 0 ≤ᵐ[μ] f) : + ennreal.of_real (∫ x, f x ∂μ) = ∫⁻ x, ennreal.of_real (f x) ∂μ := +begin + simp_rw [integral_congr_ae (show f =ᵐ[μ] λ x, ∥f x∥, + by { filter_upwards [f_nn] with x hx, + rw [real.norm_eq_abs, abs_eq_self.mpr hx], }), + of_real_integral_norm_eq_lintegral_nnnorm hfi, ←of_real_norm_eq_coe_nnnorm], + apply lintegral_congr_ae, + filter_upwards [f_nn] with x hx, + exact congr_arg ennreal.of_real (by rw [real.norm_eq_abs, abs_eq_self.mpr hx]), +end + lemma integral_to_real {f : α → ℝ≥0∞} (hfm : ae_measurable f μ) (hf : ∀ᵐ x ∂μ, f x < ∞) : ∫ a, (f a).to_real ∂μ = (∫⁻ a, f a ∂μ).to_real := begin From 785f62c695ccb62810a794268fa250f78e06c94a Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 2 May 2022 17:28:47 +0000 Subject: [PATCH 022/153] feat(algebra/star/prod): elementwise `star` operator (#13856) The lemmas and instances this provides are inspired by `algebra/star/pi`, and appear in the same order. We should have these instances anyway for completness, but the motivation is to make it easy to talk about the continuity of `star` on `units R` via the `units.embed_product_star` lemma. --- src/algebra/star/prod.lean | 52 ++++++++++++++++++++++++++++++++++++++ 1 file changed, 52 insertions(+) create mode 100644 src/algebra/star/prod.lean diff --git a/src/algebra/star/prod.lean b/src/algebra/star/prod.lean new file mode 100644 index 0000000000000..888a46e9bb9c7 --- /dev/null +++ b/src/algebra/star/prod.lean @@ -0,0 +1,52 @@ +/- +Copyright (c) 2022 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import algebra.star.basic +import algebra.ring.prod +import algebra.module.prod + +/-! +# `star` on product types + +We put a `has_star` structure on product types that operates elementwise. +-/ + +universes u v w +variables {R : Type u} {S : Type v} + +namespace prod + +instance [has_star R] [has_star S] : has_star (R × S) := +{ star := λ x, (star x.1, star x.2) } + +@[simp] lemma fst_star [has_star R] [has_star S] (x : R × S) : (star x).1 = star x.1 := rfl +@[simp] lemma snd_star [has_star R] [has_star S] (x : R × S) : (star x).2 = star x.2 := rfl + +lemma star_def [has_star R] [has_star S] (x : R × S) : star x = (star x.1, star x.2) := rfl + +instance [has_involutive_star R] [has_involutive_star S] : has_involutive_star (R × S) := +{ star_involutive := λ _, prod.ext (star_star _) (star_star _) } + +instance [semigroup R] [semigroup S] [star_semigroup R] [star_semigroup S] : + star_semigroup (R × S) := +{ star_mul := λ _ _, prod.ext (star_mul _ _) (star_mul _ _) } + +instance [add_monoid R] [add_monoid S] [star_add_monoid R] [star_add_monoid S] : + star_add_monoid (R × S) := +{ star_add := λ _ _, prod.ext (star_add _ _) (star_add _ _) } + +instance [non_unital_semiring R] [non_unital_semiring S] [star_ring R] [star_ring S] : + star_ring (R × S) := +{ ..prod.star_add_monoid, ..(prod.star_semigroup : star_semigroup (R × S)) } + +instance {α : Type w} [has_scalar α R] [has_scalar α S] [has_star α] [has_star R] [has_star S] + [star_module α R] [star_module α S] : + star_module α (R × S) := +{ star_smul := λ r x, prod.ext (star_smul _ _) (star_smul _ _) } + +end prod + +@[simp] lemma units.embed_product_star [monoid R] [star_semigroup R] (u : Rˣ) : + units.embed_product R (star u) = star (units.embed_product R u) := rfl From 3e2f214527310a1003a340f2901324ee2d4389b1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 2 May 2022 17:28:48 +0000 Subject: [PATCH 023/153] feat(logic/basic): Generalize `congr_fun_heq` (#13867) The lemma holds for arbitrary heterogeneous equalities, not only that given by casts. --- src/logic/basic.lean | 8 ++++---- src/set_theory/game/ordinal.lean | 2 +- 2 files changed, 5 insertions(+), 5 deletions(-) diff --git a/src/logic/basic.lean b/src/logic/basic.lean index fe3d02aa4ae8c..4eb9f388c4c4c 100644 --- a/src/logic/basic.lean +++ b/src/logic/basic.lean @@ -152,6 +152,10 @@ assume ⟨h⟩, h.elim @[simp] theorem exists_pempty {P : pempty → Prop} : (∃ x : pempty, P x) ↔ false := ⟨λ h, by { cases h with w, cases w }, false.elim⟩ +lemma congr_heq {α β γ : Sort*} {f : α → γ} {g : β → γ} {x : α} {y : β} (h₁ : f == g) + (h₂ : x == y) : f x = g y := +by { cases h₂, cases h₁, refl } + lemma congr_arg_heq {α} {β : α → Sort*} (f : ∀ a, β a) : ∀ {a₁ a₂ : α}, a₁ = a₂ → f a₁ == f a₂ | a _ rfl := heq.rfl @@ -885,10 +889,6 @@ lemma heq_of_cast_eq : ∀ {α β : Sort*} {a : α} {a' : β} (e : α = β) (h₂ : cast e a = a'), a == a' | α ._ a a' rfl h := eq.rec_on h (heq.refl _) -lemma congr_fun_heq {α β γ : Sort*} {f : α → γ} {g : β → γ} (h₁ : β = α) (h₂ : f == g) (x : β) : - f (cast h₁ x) = g x := -by { subst h₁, rw [eq_of_heq h₂, cast_eq] } - lemma cast_eq_iff_heq {α β : Sort*} {a : α} {a' : β} {e : α = β} : cast e a = a' ↔ a == a' := ⟨heq_of_cast_eq _, λ h, by cases h; refl⟩ diff --git a/src/set_theory/game/ordinal.lean b/src/set_theory/game/ordinal.lean index 0fc3f79904c14..4c273ff985342 100644 --- a/src/set_theory/game/ordinal.lean +++ b/src/set_theory/game/ordinal.lean @@ -63,7 +63,7 @@ by { rw to_pgame, refl } @[simp] theorem to_pgame_move_left {o : ordinal} (i : o.out.α) : o.to_pgame.move_left (to_left_moves_to_pgame i) = (typein (<) i).to_pgame := -by { rw to_left_moves_to_pgame, exact congr_fun_heq _ to_pgame_move_left_heq i } +by { rw to_left_moves_to_pgame, exact congr_heq to_pgame_move_left_heq (cast_heq _ i) } theorem to_pgame_lt {a b : ordinal} (h : a < b) : a.to_pgame < b.to_pgame := begin From 0606d7ca39d44b4c8d61fe798ce4a99305c2a4f4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 2 May 2022 17:28:49 +0000 Subject: [PATCH 024/153] feat(set_theory/game/pgame): Negative of `of_lists` (#13868) Co-authored-by: Eric Wieser --- src/set_theory/game/pgame.lean | 20 +++++++++++++++++--- 1 file changed, 17 insertions(+), 3 deletions(-) diff --git a/src/set_theory/game/pgame.lean b/src/set_theory/game/pgame.lean index 717fd1bed65b1..1514828191aaa 100644 --- a/src/set_theory/game/pgame.lean +++ b/src/set_theory/game/pgame.lean @@ -4,8 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Reid Barton, Mario Carneiro, Isabel Longbottom, Scott Morrison -/ import data.fin.basic -import data.nat.cast -import logic.embedding +import data.list.basic /-! # Combinatorial (pre-)games. @@ -145,7 +144,7 @@ Construct a pre-game from list of pre-games describing the available moves for L -- TODO define this at the level of games, as well, and perhaps also for finsets of games. def of_lists (L R : list pgame.{u}) : pgame.{u} := mk (ulift (fin L.length)) (ulift (fin R.length)) - (λ i, L.nth_le i.down i.down.is_lt) (λ j, R.nth_le j.down.val j.down.is_lt) + (λ i, L.nth_le i.down i.down.is_lt) (λ j, R.nth_le j.down j.down.prop) lemma left_moves_of_lists (L R : list pgame) : (of_lists L R).left_moves = ulift (fin L.length) := rfl @@ -704,6 +703,21 @@ begin congr; funext i; cases i end +@[simp] lemma neg_of_lists (L R : list pgame) : + -of_lists L R = of_lists (R.map (λ x, -x)) (L.map (λ x, -x)) := +begin + simp only [of_lists, neg_def, list.length_map, list.nth_le_map', eq_self_iff_true, true_and], + split, all_goals + { apply hfunext, + { simp }, + { intros a a' ha, + congr' 2, + have : ∀ {m n} (h₁ : m = n) {b : ulift (fin m)} {c : ulift (fin n)} (h₂ : b == c), + (b.down : ℕ) = ↑c.down, + { rintros m n rfl b c rfl, refl }, + exact this (list.length_map _ _).symm ha } } +end + /-- An explicit equivalence between the moves for Left in `-x` and the moves for Right in `x`. -/ -- This equivalence is useful to avoid having to use `cases` unnecessarily. def left_moves_neg (x : pgame) : (-x).left_moves ≃ x.right_moves := From aa921ef3bb68c9cf18683dcd6eb737c794184b23 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Mon, 2 May 2022 17:28:50 +0000 Subject: [PATCH 025/153] docs(set_theory/game/pgame): Fix note on `pgame` (#13880) We never actually quotient by extensionality. What we quotient by is game equivalence. --- src/set_theory/game/pgame.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/set_theory/game/pgame.lean b/src/set_theory/game/pgame.lean index 1514828191aaa..85b789eb21f05 100644 --- a/src/set_theory/game/pgame.lean +++ b/src/set_theory/game/pgame.lean @@ -108,7 +108,7 @@ open function universes u /-- The type of pre-games, before we have quotiented - by extensionality. In ZFC, a combinatorial game is constructed from + by equivalence (`pgame.setoid`). In ZFC, a combinatorial game is constructed from two sets of combinatorial games that have been constructed at an earlier stage. To do this in type theory, we say that a pre-game is built inductively from two families of pre-games indexed over any type From 2b0aeda1dbee991cd82ca3fd3d218b46e3d51878 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Mon, 2 May 2022 19:32:04 +0000 Subject: [PATCH 026/153] feat(measure/function/l*_space): a sample of useful lemmas on L^p spaces (#13823) Used in #13690 --- src/measure_theory/function/l1_space.lean | 49 +++++++++--- src/measure_theory/function/l2_space.lean | 30 ++++++++ src/measure_theory/function/lp_space.lean | 76 +++++++++++++++---- src/measure_theory/integral/bochner.lean | 19 ++++- src/measure_theory/measure/ae_measurable.lean | 17 +++++ src/measure_theory/measure/measure_space.lean | 9 +++ src/probability/integration.lean | 8 +- 7 files changed, 180 insertions(+), 28 deletions(-) diff --git a/src/measure_theory/function/l1_space.lean b/src/measure_theory/function/l1_space.lean index 8bf4ee4b080ca..807aed7299317 100644 --- a/src/measure_theory/function/l1_space.lean +++ b/src/measure_theory/function/l1_space.lean @@ -446,6 +446,25 @@ end lemma integrable_const [is_finite_measure μ] (c : β) : integrable (λ x : α, c) μ := integrable_const_iff.2 $ or.inr $ measure_lt_top _ _ +lemma mem_ℒp.integrable_norm_rpow {f : α → β} {p : ℝ≥0∞} + (hf : mem_ℒp f p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) : + integrable (λ (x : α), ∥f x∥ ^ p.to_real) μ := +begin + rw ← mem_ℒp_one_iff_integrable, + exact hf.norm_rpow hp_ne_zero hp_ne_top, +end + +lemma mem_ℒp.integrable_norm_rpow' [is_finite_measure μ] {f : α → β} {p : ℝ≥0∞} + (hf : mem_ℒp f p μ) : + integrable (λ (x : α), ∥f x∥ ^ p.to_real) μ := +begin + by_cases h_zero : p = 0, + { simp [h_zero, integrable_const] }, + by_cases h_top : p = ∞, + { simp [h_top, integrable_const] }, + exact hf.integrable_norm_rpow h_zero h_top +end + lemma integrable.mono_measure {f : α → β} (h : integrable f ν) (hμ : μ ≤ ν) : integrable f μ := ⟨h.ae_strongly_measurable.mono_measure hμ, h.has_finite_integral.mono_measure hμ⟩ @@ -552,13 +571,14 @@ lemma integrable.add {f g : α → β} (hf : integrable f μ) (hg : integrable g μ) : integrable (f + g) μ := ⟨hf.ae_strongly_measurable.add hg.ae_strongly_measurable, hf.add' hg⟩ +lemma integrable_finset_sum' {ι} (s : finset ι) + {f : ι → α → β} (hf : ∀ i ∈ s, integrable (f i) μ) : integrable (∑ i in s, f i) μ := +finset.sum_induction f (λ g, integrable g μ) (λ _ _, integrable.add) + (integrable_zero _ _ _) hf + lemma integrable_finset_sum {ι} (s : finset ι) {f : ι → α → β} (hf : ∀ i ∈ s, integrable (f i) μ) : integrable (λ a, ∑ i in s, f i a) μ := -begin - simp only [← finset.sum_apply], - exact finset.sum_induction f (λ g, integrable g μ) (λ _ _, integrable.add) - (integrable_zero _ _ _) hf, -end +by simpa only [← finset.sum_apply] using integrable_finset_sum' s hf lemma integrable.neg {f : α → β} (hf : integrable f μ) : integrable (-f) μ := ⟨hf.ae_strongly_measurable.neg, hf.has_finite_integral.neg⟩ @@ -584,6 +604,10 @@ lemma integrable.norm {f : α → β} (hf : integrable f μ) : integrable (λa, ∥f a∥) μ := ⟨hf.ae_strongly_measurable.norm, hf.has_finite_integral.norm⟩ +lemma integrable.abs {f : α → ℝ} (hf : integrable f μ) : + integrable (λa, |f a|) μ := +by simpa [← real.norm_eq_abs] using hf.norm + lemma integrable_norm_iff {f : α → β} (hf : ae_strongly_measurable f μ) : integrable (λa, ∥f a∥) μ ↔ integrable f μ := by simp_rw [integrable, and_iff_right hf, and_iff_right hf.norm, has_finite_integral_norm_iff] @@ -804,13 +828,12 @@ mem_ℒp_one_iff_integrable.1 $ mem_ℒ1_to_real_of_lintegral_ne_top hfm hfi section pos_part /-! ### Lemmas used for defining the positive part of a `L¹` function -/ -lemma integrable.max_zero {f : α → ℝ} (hf : integrable f μ) : integrable (λ a, max (f a) 0) μ := +lemma integrable.pos_part {f : α → ℝ} (hf : integrable f μ) : integrable (λ a, max (f a) 0) μ := ⟨(hf.ae_strongly_measurable.ae_measurable.max ae_measurable_const).ae_strongly_measurable, hf.has_finite_integral.max_zero⟩ -lemma integrable.min_zero {f : α → ℝ} (hf : integrable f μ) : integrable (λ a, min (f a) 0) μ := -⟨(hf.ae_strongly_measurable.ae_measurable.min ae_measurable_const).ae_strongly_measurable, - hf.has_finite_integral.min_zero⟩ +lemma integrable.neg_part {f : α → ℝ} (hf : integrable f μ) : integrable (λ a, max (-f a) 0) μ := +hf.neg.pos_part end pos_part @@ -829,10 +852,18 @@ lemma integrable.const_mul {f : α → ℝ} (h : integrable f μ) (c : ℝ) : integrable (λ x, c * f x) μ := integrable.smul c h +lemma integrable.const_mul' {f : α → ℝ} (h : integrable f μ) (c : ℝ) : + integrable ((λ (x : α), c) * f) μ := +integrable.smul c h + lemma integrable.mul_const {f : α → ℝ} (h : integrable f μ) (c : ℝ) : integrable (λ x, f x * c) μ := by simp_rw [mul_comm, h.const_mul _] +lemma integrable.mul_const' {f : α → ℝ} (h : integrable f μ) (c : ℝ) : + integrable (f * (λ (x : α), c)) μ := +integrable.mul_const h c + lemma integrable.div_const {f : α → ℝ} (h : integrable f μ) (c : ℝ) : integrable (λ x, f x / c) μ := by simp_rw [div_eq_mul_inv, h.mul_const] diff --git a/src/measure_theory/function/l2_space.lean b/src/measure_theory/function/l2_space.lean index af5cc67b9f714..4c47516f2b619 100644 --- a/src/measure_theory/function/l2_space.lean +++ b/src/measure_theory/function/l2_space.lean @@ -26,11 +26,41 @@ open topological_space measure_theory measure_theory.Lp open_locale nnreal ennreal measure_theory namespace measure_theory + +section + +variables {α F : Type*} {m : measurable_space α} {μ : measure α} [normed_group F] + +lemma mem_ℒp.integrable_sq {f : α → ℝ} (h : mem_ℒp f 2 μ) : + integrable (λ x, (f x)^2) μ := +by simpa [real.norm_eq_abs, ← mem_ℒp_one_iff_integrable] + using h.norm_rpow ennreal.two_ne_zero ennreal.two_ne_top + +lemma mem_ℒp_two_iff_integrable_sq_norm {f : α → F} (hf : ae_strongly_measurable f μ) : + mem_ℒp f 2 μ ↔ integrable (λ x, ∥f x∥^2) μ := +begin + rw ← mem_ℒp_one_iff_integrable, + convert (mem_ℒp_norm_rpow_iff hf ennreal.two_ne_zero ennreal.two_ne_top).symm, + { simp }, + { rw [div_eq_mul_inv, ennreal.mul_inv_cancel ennreal.two_ne_zero ennreal.two_ne_top] } +end + +lemma mem_ℒp_two_iff_integrable_sq {f : α → ℝ} (hf : ae_strongly_measurable f μ) : + mem_ℒp f 2 μ ↔ integrable (λ x, (f x)^2) μ := +begin + convert mem_ℒp_two_iff_integrable_sq_norm hf, + ext x, + simp [real.norm_eq_abs], +end + +end + namespace L2 variables {α E F 𝕜 : Type*} [is_R_or_C 𝕜] [measurable_space α] {μ : measure α} [inner_product_space 𝕜 E] [normed_group F] + local notation `⟪`x`, `y`⟫` := @inner 𝕜 _ _ x y lemma snorm_rpow_two_norm_lt_top (f : Lp F 2 μ) : snorm (λ x, ∥f x∥ ^ (2 : ℝ)) 1 μ < ∞ := diff --git a/src/measure_theory/function/lp_space.lean b/src/measure_theory/function/lp_space.lean index ca7ad8f4d0411..425c5a936b3fa 100644 --- a/src/measure_theory/function/lp_space.lean +++ b/src/measure_theory/function/lp_space.lean @@ -131,7 +131,8 @@ by simp_rw [snorm_eq_lintegral_rpow_nnnorm one_ne_zero ennreal.coe_ne_top, ennre /-- The property that `f:α→E` is ae strongly measurable and `(∫ ∥f a∥^p ∂μ)^(1/p)` is finite if `p < ∞`, or `ess_sup f < ∞` if `p = ∞`. -/ -def mem_ℒp {α} {m : measurable_space α} (f : α → E) (p : ℝ≥0∞) (μ : measure α) : Prop := +def mem_ℒp {α} {m : measurable_space α} + (f : α → E) (p : ℝ≥0∞) (μ : measure α . volume_tac) : Prop := ae_strongly_measurable f μ ∧ snorm f p μ < ∞ lemma mem_ℒp.ae_strongly_measurable {f : α → E} {p : ℝ≥0∞} (h : mem_ℒp f p μ) : @@ -339,6 +340,15 @@ begin simp, end +lemma mem_ℒp_top_const (c : E) : mem_ℒp (λ a:α, c) ∞ μ := +begin + refine ⟨ae_strongly_measurable_const, _⟩, + by_cases h : μ = 0, + { simp only [h, snorm_measure_zero, with_top.zero_lt_top] }, + { rw snorm_const _ ennreal.top_ne_zero h, + simp only [ennreal.top_to_real, div_zero, ennreal.rpow_zero, mul_one, ennreal.coe_lt_top] } +end + lemma mem_ℒp_const_iff {p : ℝ≥0∞} {c : E} (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) : mem_ℒp (λ x : α, c) p μ ↔ c = 0 ∨ μ set.univ < ∞ := begin @@ -1141,6 +1151,14 @@ begin exact (hf i (s.mem_insert_self i)).add (ih (λ j hj, hf j (finset.mem_insert_of_mem hj))), }, end +lemma mem_ℒp_finset_sum' {ι} (s : finset ι) {f : ι → α → E} (hf : ∀ i ∈ s, mem_ℒp (f i) p μ) : + mem_ℒp (∑ i in s, f i) p μ := +begin + convert mem_ℒp_finset_sum s hf, + ext x, + simp, +end + end has_measurable_add end borel_space @@ -1324,7 +1342,7 @@ by simp [hfp.2] /-- Lp space -/ def Lp {α} (E : Type*) {m : measurable_space α} [normed_group E] - (p : ℝ≥0∞) (μ : measure α) : add_subgroup (α →ₘ[μ] E) := + (p : ℝ≥0∞) (μ : measure α . volume_tac) : add_subgroup (α →ₘ[μ] E) := { carrier := {f | snorm f p μ < ∞}, zero_mem' := by simp [snorm_congr_ae ae_eq_fun.coe_fn_zero, snorm_zero], add_mem' := λ f g hf hg, by simp [snorm_congr_ae (ae_eq_fun.coe_fn_add _ _), @@ -1820,21 +1838,45 @@ end end indicator_const_Lp +lemma mem_ℒp.norm_rpow_div {f : α → E} + (hf : mem_ℒp f p μ) (q : ℝ≥0∞) : + mem_ℒp (λ (x : α), ∥f x∥ ^ q.to_real) (p/q) μ := +begin + refine ⟨(hf.1.norm.ae_measurable.pow_const q.to_real).ae_strongly_measurable, _⟩, + by_cases q_top : q = ∞, { simp [q_top] }, + by_cases q_zero : q = 0, + { simp [q_zero], + by_cases p_zero : p = 0, { simp [p_zero] }, + rw ennreal.div_zero p_zero, + exact (mem_ℒp_top_const (1 : ℝ)).2 }, + rw snorm_norm_rpow _ (ennreal.to_real_pos q_zero q_top), + apply ennreal.rpow_lt_top_of_nonneg ennreal.to_real_nonneg, + rw [ennreal.of_real_to_real q_top, div_eq_mul_inv, mul_assoc, + ennreal.inv_mul_cancel q_zero q_top, mul_one], + exact hf.2.ne +end + +lemma mem_ℒp_norm_rpow_iff {q : ℝ≥0∞} {f : α → E} (hf : ae_strongly_measurable f μ) + (q_zero : q ≠ 0) (q_top : q ≠ ∞) : + mem_ℒp (λ (x : α), ∥f x∥ ^ q.to_real) (p/q) μ ↔ mem_ℒp f p μ := +begin + refine ⟨λ h, _, λ h, h.norm_rpow_div q⟩, + apply (mem_ℒp_norm_iff hf).1, + convert h.norm_rpow_div (q⁻¹), + { ext x, + rw [real.norm_eq_abs, real.abs_rpow_of_nonneg (norm_nonneg _), ← real.rpow_mul (abs_nonneg _), + ennreal.to_real_inv, mul_inv_cancel, abs_of_nonneg (norm_nonneg _), real.rpow_one], + simp [ennreal.to_real_eq_zero_iff, not_or_distrib, q_zero, q_top] }, + { rw [div_eq_mul_inv, inv_inv, div_eq_mul_inv, mul_assoc, ennreal.inv_mul_cancel q_zero q_top, + mul_one] } +end + lemma mem_ℒp.norm_rpow {f : α → E} (hf : mem_ℒp f p μ) (hp_ne_zero : p ≠ 0) (hp_ne_top : p ≠ ∞) : mem_ℒp (λ (x : α), ∥f x∥ ^ p.to_real) 1 μ := begin - refine ⟨(hf.1.norm.ae_measurable.pow_const p.to_real).ae_strongly_measurable, _⟩, - have := hf.snorm_ne_top, - rw snorm_eq_lintegral_rpow_nnnorm hp_ne_zero hp_ne_top at this, - rw snorm_one_eq_lintegral_nnnorm, - convert ennreal.rpow_lt_top_of_nonneg (@ennreal.to_real_nonneg p) this, - rw [← ennreal.rpow_mul, one_div_mul_cancel (ennreal.to_real_pos hp_ne_zero hp_ne_top).ne.symm, - ennreal.rpow_one], - congr, - ext1 x, - rw [ennreal.coe_rpow_of_nonneg _ ennreal.to_real_nonneg, real.nnnorm_of_nonneg], - congr + convert hf.norm_rpow_div p, + rw [div_eq_mul_inv, ennreal.mul_inv_cancel hp_ne_zero hp_ne_top], end end measure_theory @@ -2090,6 +2132,14 @@ section pos_part lemma lipschitz_with_pos_part : lipschitz_with 1 (λ (x : ℝ), max x 0) := lipschitz_with.of_dist_le_mul $ λ x y, by simp [dist, abs_max_sub_max_le_abs] +lemma _root_.measure_theory.mem_ℒp.pos_part {f : α → ℝ} (hf : mem_ℒp f p μ) : + mem_ℒp (λ x, max (f x) 0) p μ := +lipschitz_with_pos_part.comp_mem_ℒp (max_eq_right le_rfl) hf + +lemma _root_.measure_theory.mem_ℒp.neg_part {f : α → ℝ} (hf : mem_ℒp f p μ) : + mem_ℒp (λ x, max (-f x) 0) p μ := +lipschitz_with_pos_part.comp_mem_ℒp (max_eq_right le_rfl) hf.neg + /-- Positive part of a function in `L^p`. -/ def pos_part (f : Lp ℝ p μ) : Lp ℝ p μ := lipschitz_with_pos_part.comp_Lp (max_eq_right le_rfl) f diff --git a/src/measure_theory/integral/bochner.lean b/src/measure_theory/integral/bochner.lean index 2d76699bfd7f1..6c728afa447d5 100644 --- a/src/measure_theory/integral/bochner.lean +++ b/src/measure_theory/integral/bochner.lean @@ -573,8 +573,8 @@ begin have := (to_simple_func f).pos_part_sub_neg_part, conv_lhs {rw ← this}, refl, }, - { exact (simple_func.integrable f).max_zero.congr ae_eq₁ }, - { exact (simple_func.integrable f).neg.max_zero.congr ae_eq₂ } + { exact (simple_func.integrable f).pos_part.congr ae_eq₁ }, + { exact (simple_func.integrable f).neg_part.congr ae_eq₂ } end end pos_part @@ -1105,6 +1105,21 @@ begin simp [ha] end +lemma mem_ℒp.snorm_eq_integral_rpow_norm {f : α → H} {p : ℝ≥0∞} (hp1 : p ≠ 0) (hp2 : p ≠ ∞) + (hf : mem_ℒp f p μ) : + snorm f p μ = ennreal.of_real ((∫ a, ∥f a∥ ^ p.to_real ∂μ) ^ (p.to_real ⁻¹)) := +begin + have A : ∫⁻ (a : α), ennreal.of_real (∥f a∥ ^ p.to_real) ∂μ = ∫⁻ (a : α), ∥f a∥₊ ^ p.to_real ∂μ, + { apply lintegral_congr (λ x, _), + rw [← of_real_rpow_of_nonneg (norm_nonneg _) to_real_nonneg, of_real_norm_eq_coe_nnnorm] }, + simp only [snorm_eq_lintegral_rpow_nnnorm hp1 hp2, one_div], + rw integral_eq_lintegral_of_nonneg_ae, rotate, + { exact eventually_of_forall (λ x, real.rpow_nonneg_of_nonneg (norm_nonneg _) _) }, + { exact (hf.ae_strongly_measurable.norm.ae_measurable.pow_const _).ae_strongly_measurable }, + rw [A, ← of_real_rpow_of_nonneg to_real_nonneg (inv_nonneg.2 to_real_nonneg), of_real_to_real], + exact (lintegral_rpow_nnnorm_lt_top_of_snorm_lt_top hp1 hp2 hf.2).ne +end + end normed_group lemma integral_mono_ae {f g : α → ℝ} (hf : integrable f μ) (hg : integrable g μ) (h : f ≤ᵐ[μ] g) : diff --git a/src/measure_theory/measure/ae_measurable.lean b/src/measure_theory/measure/ae_measurable.lean index 4303872fee6ef..13816f0ef631a 100644 --- a/src/measure_theory/measure/ae_measurable.lean +++ b/src/measure_theory/measure/ae_measurable.lean @@ -133,6 +133,23 @@ lemma comp_measurable' {ν : measure δ} {f : α → δ} {g : δ → β} (hg : a (hf : measurable f) (h : μ.map f ≪ ν) : ae_measurable (g ∘ f) μ := (hg.mono' h).comp_measurable hf +lemma map_map_of_ae_measurable {g : β → γ} {f : α → β} + (hg : ae_measurable g (measure.map f μ)) (hf : ae_measurable f μ) : + (μ.map f).map g = μ.map (g ∘ f) := +begin + ext1 s hs, + let g' := hg.mk g, + have A : map g (map f μ) = map g' (map f μ), + { apply measure_theory.measure.map_congr, + exact hg.ae_eq_mk }, + have B : map (g ∘ f) μ = map (g' ∘ f) μ, + { apply measure_theory.measure.map_congr, + exact ae_of_ae_map hf hg.ae_eq_mk }, + simp only [A, B, hs, hg.measurable_mk.ae_measurable.comp_ae_measurable hf, hg.measurable_mk, + hg.measurable_mk hs, hf, map_apply, map_apply_of_ae_measurable], + refl, +end + @[measurability] lemma prod_mk {f : α → β} {g : α → γ} (hf : ae_measurable f μ) (hg : ae_measurable g μ) : ae_measurable (λ x, (f x, g x)) μ := diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index 1480575acbb02..2cf0c911f4c6a 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -2106,6 +2106,10 @@ instance restrict.is_finite_measure (μ : measure α) [hs : fact (μ s < ∞)] : lemma measure_lt_top (μ : measure α) [is_finite_measure μ] (s : set α) : μ s < ∞ := (measure_mono (subset_univ s)).trans_lt is_finite_measure.measure_univ_lt_top +instance is_finite_measure_restrict (μ : measure α) (s : set α) [h : is_finite_measure μ] : + is_finite_measure (μ.restrict s) := +⟨by simp [measure_lt_top μ s]⟩ + lemma measure_ne_top (μ : measure α) [is_finite_measure μ] (s : set α) : μ s ≠ ∞ := ne_of_lt (measure_lt_top μ s) @@ -2216,6 +2220,7 @@ include m0 class is_probability_measure (μ : measure α) : Prop := (measure_univ : μ univ = 1) export is_probability_measure (measure_univ) +attribute [simp] is_probability_measure.measure_univ @[priority 100] instance is_probability_measure.to_is_finite_measure (μ : measure α) [is_probability_measure μ] : @@ -2251,6 +2256,10 @@ begin { exact measure_ne_top _ _ } end +lemma is_probability_measure_map [is_probability_measure μ] {f : α → β} (hf : ae_measurable f μ) : + is_probability_measure (map f μ) := +⟨by simp [map_apply_of_ae_measurable, hf]⟩ + end is_probability_measure section no_atoms diff --git a/src/probability/integration.lean b/src/probability/integration.lean index d8be5cdb4012d..0fdd1e0fc7d43 100644 --- a/src/probability/integration.lean +++ b/src/probability/integration.lean @@ -210,10 +210,10 @@ begin have hm3 : ae_measurable Ym μ := hY.1.ae_measurable.neg.max ae_measurable_const, have hm4 : ae_measurable Yp μ := hY.1.ae_measurable.max ae_measurable_const, - have hv1 : integrable Xm μ := hX.neg.max_zero, - have hv2 : integrable Xp μ := hX.max_zero, - have hv3 : integrable Ym μ := hY.neg.max_zero, - have hv4 : integrable Yp μ := hY.max_zero, + have hv1 : integrable Xm μ := hX.neg_part, + have hv2 : integrable Xp μ := hX.pos_part, + have hv3 : integrable Ym μ := hY.neg_part, + have hv4 : integrable Yp μ := hY.pos_part, have hi1 : indep_fun Xm Ym μ := hXY.comp negm negm, have hi2 : indep_fun Xp Ym μ := hXY.comp posm negm, From c44091fcf8c4a87d411ed6c7b86a325b4ff1532c Mon Sep 17 00:00:00 2001 From: tb65536 Date: Mon, 2 May 2022 19:32:06 +0000 Subject: [PATCH 027/153] feat(data/zmod/basic): Generalize `zmod.card` (#13887) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR generalizes `zmod.card` from assuming `[fact (0 < n)]` to assuming `[fintype (zmod n)]`. Note that the latter was already part of the statement, but was previously deduced from the instance `instance fintype : Π (n : ℕ) [fact (0 < n)], fintype (zmod n)` on line 80. --- src/data/zmod/basic.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/src/data/zmod/basic.lean b/src/data/zmod/basic.lean index c083d660d6e79..accb4e6b89a01 100644 --- a/src/data/zmod/basic.lean +++ b/src/data/zmod/basic.lean @@ -84,11 +84,11 @@ instance fintype : Π (n : ℕ) [fact (0 < n)], fintype (zmod n) instance infinite : infinite (zmod 0) := int.infinite -@[simp] lemma card (n : ℕ) [fact (0 < n)] : fintype.card (zmod n) = n := +@[simp] lemma card (n : ℕ) [fintype (zmod n)] : fintype.card (zmod n) = n := begin casesI n, - { exfalso, exact nat.not_lt_zero 0 (fact.out _) }, - { exact fintype.card_fin (n+1) } + { exfalso, exact not_fintype (zmod 0) }, + { convert fintype.card_fin (n+1) } end instance decidable_eq : Π (n : ℕ), decidable_eq (zmod n) From 174120760f712153fcca4a7d603f0236f76ff0e6 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 2 May 2022 20:22:14 +0000 Subject: [PATCH 028/153] =?UTF-8?q?feat(analysis/normed=5Fspace/exponentia?= =?UTF-8?q?l):=20`Ae=E1=B4=AE=20=3D=20e=E1=B4=AEA`=20if=20`AB=20=3D=20BA`?= =?UTF-8?q?=20(#13881)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This commit shows that the exponenential commutes if the exponent does. This generalizes a previous weaker result. --- src/analysis/normed_space/exponential.lean | 18 ++++++++++++++---- src/topology/algebra/infinite_sum.lean | 9 +++++++++ 2 files changed, 23 insertions(+), 4 deletions(-) diff --git a/src/analysis/normed_space/exponential.lean b/src/analysis/normed_space/exponential.lean index 4a4a6d56b139d..75c601672c40b 100644 --- a/src/analysis/normed_space/exponential.lean +++ b/src/analysis/normed_space/exponential.lean @@ -123,6 +123,20 @@ begin simp [h] end +variables (𝕂) + +lemma commute.exp_right [t2_space 𝔸] {x y : 𝔸} (h : commute x y) : commute x (exp 𝕂 𝔸 y) := +begin + rw exp_eq_tsum, + exact commute.tsum_right x (λ n, (h.pow_right n).smul_right _), +end + +lemma commute.exp_left [t2_space 𝔸] {x y : 𝔸} (h : commute x y) : commute (exp 𝕂 𝔸 x) y := +(h.symm.exp_right 𝕂).symm + +lemma commute.exp [t2_space 𝔸] {x y : 𝔸} (h : commute x y) : commute (exp 𝕂 𝔸 x) (exp 𝕂 𝔸 y) := +(h.exp_left _).exp_right _ + end topological_algebra section normed @@ -421,10 +435,6 @@ begin exact ring.inverse_invertible _, end -lemma commute.exp {x y : 𝔸} (h : commute x y) : - commute (exp 𝕂 𝔸 x) (exp 𝕂 𝔸 y) := -(exp_add_of_commute h).symm.trans $ (congr_arg _ $ add_comm _ _).trans (exp_add_of_commute h.symm) - end /-- In a Banach-algebra `𝔸` over `𝕂 = ℝ` or `𝕂 = ℂ`, if a family of elements `f i` mutually diff --git a/src/topology/algebra/infinite_sum.lean b/src/topology/algebra/infinite_sum.lean index dfedbec0240aa..78292bb753a0e 100644 --- a/src/topology/algebra/infinite_sum.lean +++ b/src/topology/algebra/infinite_sum.lean @@ -777,6 +777,15 @@ lemma summable.tsum_mul_left (a) (hf : summable f) : ∑'b, a * f b = a * ∑'b, lemma summable.tsum_mul_right (a) (hf : summable f) : (∑'b, f b * a) = (∑'b, f b) * a := (hf.has_sum.mul_right _).tsum_eq +lemma commute.tsum_right (a) (h : ∀ b, commute a (f b)) : commute a (∑' b, f b) := +if hf : summable f then + (hf.tsum_mul_left a).symm.trans ((congr_arg _ $ funext h).trans (hf.tsum_mul_right a)) +else + (tsum_eq_zero_of_not_summable hf).symm ▸ commute.zero_right _ + +lemma commute.tsum_left (a) (h : ∀ b, commute (f b) a) : commute (∑' b, f b) a := +(commute.tsum_right _ $ λ b, (h b).symm).symm + end tsum end topological_semiring From ca1551c26a375d70da9a2c88f99ee650c596da68 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 2 May 2022 21:37:01 +0000 Subject: [PATCH 029/153] feat(data/finset/n_ary): Binary image of finsets (#13718) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define `finset.image₂`, the binary map of finsets. Golf `data.finset.pointwise` using it. --- src/data/finset/n_ary.lean | 245 +++++++++++++++++++++++++++++++++ src/data/finset/pointwise.lean | 244 ++++++++++---------------------- src/order/filter/n_ary.lean | 3 +- 3 files changed, 321 insertions(+), 171 deletions(-) create mode 100644 src/data/finset/n_ary.lean diff --git a/src/data/finset/n_ary.lean b/src/data/finset/n_ary.lean new file mode 100644 index 0000000000000..79e9f331d4769 --- /dev/null +++ b/src/data/finset/n_ary.lean @@ -0,0 +1,245 @@ +/- +Copyright (c) 2022 Yaël Dillies. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies +-/ +import data.finset.prod + +/-! +# N-ary images of finsets + +This file defines `finset.image₂`, the binary image of finsets. This is the finset version of +`set.image2`. This is mostly useful to define pointwise operations. + +## Notes + +This file is very similar to the n-ary section of `data.set.basic` and to `order.filter.n_ary`. +Please keep them in sync. + +We do not define `finset.image₃` as its only purpose would be to prove properties of `finset.image₂` +and `set.image2` already fulfills this task. +-/ + +open function set + +namespace finset + +variables {α α' β β' γ γ' δ δ' ε ε' : Type*} + [decidable_eq α'] [decidable_eq β'] [decidable_eq γ] [decidable_eq δ] [decidable_eq δ'] + [decidable_eq ε] [decidable_eq ε'] + {f f' : α → β → γ} {g g' : α → β → γ → δ} {s s' : finset α} {t t' : finset β} {u u' : finset γ} + {a a' : α} {b b' : β} {c : γ} + +/-- The image of a binary function `f : α → β → γ` as a function `finset α → finset β → finset γ`. +Mathematically this should be thought of as the image of the corresponding function `α × β → γ`. -/ +def image₂ (f : α → β → γ) (s : finset α) (t : finset β) : finset γ := +(s.product t).image $ uncurry f + +@[simp] lemma mem_image₂ : c ∈ image₂ f s t ↔ ∃ a b, a ∈ s ∧ b ∈ t ∧ f a b = c := +by simp [image₂, and_assoc] + +@[simp, norm_cast] lemma coe_image₂ (f : α → β → γ) (s : finset α) (t : finset β) : + (image₂ f s t : set γ) = set.image2 f s t := +set.ext $ λ _, mem_image₂ + +lemma card_image₂_le (f : α → β → γ) (s : finset α) (t : finset β) : + (image₂ f s t).card ≤ s.card * t.card := +card_image_le.trans_eq $ card_product _ _ + +lemma card_image₂ (hf : injective2 f) (s : finset α) (t : finset β) : + (image₂ f s t).card = s.card * t.card := +(card_image_of_injective _ hf.uncurry).trans $ card_product _ _ + +lemma mem_image₂_of_mem (ha : a ∈ s) (hb : b ∈ t) : f a b ∈ image₂ f s t := +mem_image₂.2 ⟨a, b, ha, hb, rfl⟩ + +lemma mem_image₂_iff (hf : injective2 f) : f a b ∈ image₂ f s t ↔ a ∈ s ∧ b ∈ t := +by rw [←mem_coe, coe_image₂, mem_image2_iff hf, mem_coe, mem_coe] + +lemma image₂_subset (hs : s ⊆ s') (ht : t ⊆ t') : image₂ f s t ⊆ image₂ f s' t' := +by { rw [←coe_subset, coe_image₂, coe_image₂], exact image2_subset hs ht } + +lemma image₂_subset_left (ht : t ⊆ t') : image₂ f s t ⊆ image₂ f s t' := image₂_subset subset.rfl ht + +lemma image₂_subset_right (hs : s ⊆ s') : image₂ f s t ⊆ image₂ f s' t := +image₂_subset hs subset.rfl + +lemma image_subset_image₂_left (hb : b ∈ t) : (λ a, f a b) '' s ⊆ image₂ f s t := +ball_image_of_ball $ λ a ha, mem_image₂_of_mem ha hb + +lemma image_subset_image₂_right (ha : a ∈ s) : f a '' t ⊆ image₂ f s t := +ball_image_of_ball $ λ b, mem_image₂_of_mem ha + +lemma forall_image₂_iff {p : γ → Prop} : (∀ z ∈ image₂ f s t, p z) ↔ ∀ (x ∈ s) (y ∈ t), p (f x y) := +by simp_rw [←mem_coe, coe_image₂, forall_image2_iff] + +@[simp] lemma image₂_subset_iff : image₂ f s t ⊆ u ↔ ∀ (x ∈ s) (y ∈ t), f x y ∈ u := +forall_image₂_iff + +@[simp] lemma image₂_nonempty_iff : (image₂ f s t).nonempty ↔ s.nonempty ∧ t.nonempty := +by { rw [←coe_nonempty, coe_image₂], exact image2_nonempty_iff } + +lemma nonempty.image₂ (hs : s.nonempty) (ht : t.nonempty) : (image₂ f s t).nonempty := +image₂_nonempty_iff.2 ⟨hs, ht⟩ + +lemma nonempty.of_image₂_left (h : (image₂ f s t).nonempty) : s.nonempty := +(image₂_nonempty_iff.1 h).1 + +lemma nonempty.of_image₂_right (h : (image₂ f s t).nonempty) : t.nonempty := +(image₂_nonempty_iff.1 h).2 + +@[simp] lemma image₂_empty_left : image₂ f ∅ t = ∅ := coe_injective $ by simp +@[simp] lemma image₂_empty_right : image₂ f s ∅ = ∅ := coe_injective $ by simp +@[simp] lemma image₂_eq_empty_iff : image₂ f s t = ∅ ↔ s = ∅ ∨ t = ∅ := +by simp_rw [←not_nonempty_iff_eq_empty, image₂_nonempty_iff, not_and_distrib] + +@[simp] lemma image₂_singleton_left : image₂ f {a} t = t.image (f a) := ext $ λ x, by simp +@[simp] lemma image₂_singleton_right : image₂ f s {b} = s.image (λ a, f a b) := ext $ λ x, by simp + +lemma image₂_singleton : image₂ f {a} {b} = {f a b} := by simp + +lemma image₂_union_left [decidable_eq α] : image₂ f (s ∪ s') t = image₂ f s t ∪ image₂ f s' t := +coe_injective $ by { push_cast, exact image2_union_left } + +lemma image₂_union_right [decidable_eq β] : image₂ f s (t ∪ t') = image₂ f s t ∪ image₂ f s t' := +coe_injective $ by { push_cast, exact image2_union_right } + +lemma image₂_inter_subset_left [decidable_eq α] : + image₂ f (s ∩ s') t ⊆ image₂ f s t ∩ image₂ f s' t := +coe_subset.1 $ by { push_cast, exact image2_inter_subset_left } + +lemma image₂_inter_subset_right [decidable_eq β] : + image₂ f s (t ∩ t') ⊆ image₂ f s t ∩ image₂ f s t' := +coe_subset.1 $ by { push_cast, exact image2_inter_subset_right } + +lemma image₂_congr (h : ∀ (a ∈ s) (b ∈ t), f a b = f' a b) : image₂ f s t = image₂ f' s t := +coe_injective $ by { push_cast, exact image2_congr h } + +/-- A common special case of `image₂_congr` -/ +lemma image₂_congr' (h : ∀ a b, f a b = f' a b) : image₂ f s t = image₂ f' s t := +image₂_congr $ λ a _ b _, h a b + +lemma subset_image₂ {s : set α} {t : set β} (hu : ↑u ⊆ image2 f s t) : + ∃ (s' : finset α) (t' : finset β), ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ image₂ f s' t' := +begin + apply finset.induction_on' u, + { exact ⟨∅, ∅, set.empty_subset _, set.empty_subset _, empty_subset _⟩ }, + rintro a u ha _ _ ⟨s', t', hs, hs', h⟩, + obtain ⟨x, y, hx, hy, ha⟩ := hu ha, + haveI := classical.dec_eq α, + haveI := classical.dec_eq β, + refine ⟨insert x s', insert y t', _⟩, + simp_rw [coe_insert, set.insert_subset], + exact ⟨⟨hx, hs⟩, ⟨hy, hs'⟩, insert_subset.2 ⟨mem_image₂.2 ⟨x, y, mem_insert_self _ _, + mem_insert_self _ _, ha⟩, h.trans $ image₂_subset (subset_insert _ _) $ subset_insert _ _⟩⟩, +end + +/-! +### Algebraic replacement rules + +A collection of lemmas to transfer associativity, commutativity, distributivity, ... of operations +to the associativity, commutativity, distributivity, ... of `finset.image₂` of those operations. + +The proof pattern is `image₂_lemma operation_lemma`. For example, `image₂_comm mul_comm` proves that +`image₂ (*) f g = image₂ (*) g f` in a `comm_semigroup`. +-/ + +lemma image_image₂ (f : α → β → γ) (g : γ → δ) : + (image₂ f s t).image g = image₂ (λ a b, g (f a b)) s t := +coe_injective $ by { push_cast, exact image_image2 _ _ } + +lemma image₂_image_left (f : γ → β → δ) (g : α → γ) : + image₂ f (s.image g) t = image₂ (λ a b, f (g a) b) s t := +coe_injective $ by { push_cast, exact image2_image_left _ _ } + +lemma image₂_image_right (f : α → γ → δ) (g : β → γ) : + image₂ f s (t.image g) = image₂ (λ a b, f a (g b)) s t := +coe_injective $ by { push_cast, exact image2_image_right _ _ } + +lemma image₂_swap (f : α → β → γ) (s : finset α) (t : finset β) : + image₂ f s t = image₂ (λ a b, f b a) t s := +coe_injective $ by { push_cast, exact image2_swap _ _ _ } + +@[simp] lemma image₂_left [decidable_eq α] (h : t.nonempty) : image₂ (λ x y, x) s t = s := +coe_injective $ by { push_cast, exact image2_left h } + +@[simp] lemma image₂_right [decidable_eq β] (h : s.nonempty) : image₂ (λ x y, y) s t = t := +coe_injective $ by { push_cast, exact image2_right h } + +lemma image₂_assoc {γ : Type*} {u : finset γ} {f : δ → γ → ε} {g : α → β → δ} {f' : α → ε' → ε} + {g' : β → γ → ε'} (h_assoc : ∀ a b c, f (g a b) c = f' a (g' b c)) : + image₂ f (image₂ g s t) u = image₂ f' s (image₂ g' t u) := +coe_injective $ by { push_cast, exact image2_assoc h_assoc } + +lemma image₂_comm {g : β → α → γ} (h_comm : ∀ a b, f a b = g b a) : image₂ f s t = image₂ g t s := +(image₂_swap _ _ _).trans $ by simp_rw h_comm + +lemma image₂_left_comm {γ : Type*} {u : finset γ} {f : α → δ → ε} {g : β → γ → δ} {f' : α → γ → δ'} + {g' : β → δ' → ε} (h_left_comm : ∀ a b c, f a (g b c) = g' b (f' a c)) : + image₂ f s (image₂ g t u) = image₂ g' t (image₂ f' s u) := +coe_injective $ by { push_cast, exact image2_left_comm h_left_comm } + +lemma image₂_right_comm {γ : Type*} {u : finset γ} {f : δ → γ → ε} {g : α → β → δ} {f' : α → γ → δ'} + {g' : δ' → β → ε} (h_right_comm : ∀ a b c, f (g a b) c = g' (f' a c) b) : + image₂ f (image₂ g s t) u = image₂ g' (image₂ f' s u) t := +coe_injective $ by { push_cast, exact image2_right_comm h_right_comm } + +lemma image_image₂_distrib {g : γ → δ} {f' : α' → β' → δ} {g₁ : α → α'} {g₂ : β → β'} + (h_distrib : ∀ a b, g (f a b) = f' (g₁ a) (g₂ b)) : + (image₂ f s t).image g = image₂ f' (s.image g₁) (t.image g₂) := +coe_injective $ by { push_cast, exact image_image2_distrib h_distrib } + +/-- Symmetric of `finset.image₂_image_left_comm`. -/ +lemma image_image₂_distrib_left {g : γ → δ} {f' : α' → β → δ} {g' : α → α'} + (h_distrib : ∀ a b, g (f a b) = f' (g' a) b) : + (image₂ f s t).image g = image₂ f' (s.image g') t := +coe_injective $ by { push_cast, exact image_image2_distrib_left h_distrib } + +/-- Symmetric of `finset.image_image₂_right_comm`. -/ +lemma image_image₂_distrib_right {g : γ → δ} {f' : α → β' → δ} {g' : β → β'} + (h_distrib : ∀ a b, g (f a b) = f' a (g' b)) : + (image₂ f s t).image g = image₂ f' s (t.image g') := +coe_injective $ by { push_cast, exact image_image2_distrib_right h_distrib } + +/-- Symmetric of `finset.image_image₂_distrib_left`. -/ +lemma image₂_image_left_comm {f : α' → β → γ} {g : α → α'} {f' : α → β → δ} {g' : δ → γ} + (h_left_comm : ∀ a b, f (g a) b = g' (f' a b)) : + image₂ f (s.image g) t = (image₂ f' s t).image g' := +(image_image₂_distrib_left $ λ a b, (h_left_comm a b).symm).symm + +/-- Symmetric of `finset.image_image₂_distrib_right`. -/ +lemma image_image₂_right_comm {f : α → β' → γ} {g : β → β'} {f' : α → β → δ} {g' : δ → γ} + (h_right_comm : ∀ a b, f a (g b) = g' (f' a b)) : + image₂ f s (t.image g) = (image₂ f' s t).image g' := +(image_image₂_distrib_right $ λ a b, (h_right_comm a b).symm).symm + +lemma image_image₂_antidistrib {g : γ → δ} {f' : β' → α' → δ} {g₁ : β → β'} {g₂ : α → α'} + (h_antidistrib : ∀ a b, g (f a b) = f' (g₁ b) (g₂ a)) : + (image₂ f s t).image g = image₂ f' (t.image g₁) (s.image g₂) := +by { rw image₂_swap f, exact image_image₂_distrib (λ _ _, h_antidistrib _ _) } + +/-- Symmetric of `finset.image₂_image_left_anticomm`. -/ +lemma image_image₂_antidistrib_left {g : γ → δ} {f' : β' → α → δ} {g' : β → β'} + (h_antidistrib : ∀ a b, g (f a b) = f' (g' b) a) : + (image₂ f s t).image g = image₂ f' (t.image g') s := +coe_injective $ by { push_cast, exact image_image2_antidistrib_left h_antidistrib } + +/-- Symmetric of `finset.image_image₂_right_anticomm`. -/ +lemma image_image₂_antidistrib_right {g : γ → δ} {f' : β → α' → δ} {g' : α → α'} + (h_antidistrib : ∀ a b, g (f a b) = f' b (g' a)) : + (image₂ f s t).image g = image₂ f' t (s.image g') := +coe_injective $ by { push_cast, exact image_image2_antidistrib_right h_antidistrib } + +/-- Symmetric of `finset.image_image₂_antidistrib_left`. -/ +lemma image₂_image_left_anticomm {f : α' → β → γ} {g : α → α'} {f' : β → α → δ} {g' : δ → γ} + (h_left_anticomm : ∀ a b, f (g a) b = g' (f' b a)) : + image₂ f (s.image g) t = (image₂ f' t s).image g' := +(image_image₂_antidistrib_left $ λ a b, (h_left_anticomm b a).symm).symm + +/-- Symmetric of `finset.image_image₂_antidistrib_right`. -/ +lemma image_image₂_right_anticomm {f : α → β' → γ} {g : β → β'} {f' : β → α → δ} {g' : δ → γ} + (h_right_anticomm : ∀ a b, f a (g b) = g' (f' b a)) : + image₂ f s (t.image g) = (image₂ f' t s).image g' := +(image_image₂_antidistrib_right $ λ a b, (h_right_anticomm b a).symm).symm + +end finset diff --git a/src/data/finset/pointwise.lean b/src/data/finset/pointwise.lean index 1f4136ffafe29..ac275fb7102c7 100644 --- a/src/data/finset/pointwise.lean +++ b/src/data/finset/pointwise.lean @@ -3,6 +3,7 @@ Copyright (c) 2020 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn, Yaël Dillies -/ +import data.finset.n_ary import data.finset.preimage import data.set.pointwise @@ -133,7 +134,7 @@ variables [decidable_eq α] [has_mul α] {s s₁ s₂ t t₁ t₂ u : finset α} /-- The pointwise product of two finsets `s` and `t`: `s * t = {x * y | x ∈ s, y ∈ t}`. -/ @[to_additive "The pointwise sum of two finsets `s` and `t`: `s + t = {x + y | x ∈ s, y ∈ t}`."] -protected def has_mul : has_mul (finset α) := ⟨λ s t, (s.product t).image $ λ p : α × α, p.1 * p.2⟩ +protected def has_mul : has_mul (finset α) := ⟨image₂ (*)⟩ localized "attribute [instance] finset.has_mul finset.has_add" in pointwise @@ -144,62 +145,40 @@ lemma mul_def : s * t = (s.product t).image (λ p : α × α, p.1 * p.2) := rfl lemma image_mul_prod : (s.product t).image (λ x : α × α, x.fst * x.snd) = s * t := rfl @[to_additive] -lemma mem_mul {x : α} : x ∈ s * t ↔ ∃ y z, y ∈ s ∧ z ∈ t ∧ y * z = x := -by simp only [finset.mul_def, and.assoc, mem_image, exists_prop, prod.exists, mem_product] +lemma mem_mul {x : α} : x ∈ s * t ↔ ∃ y z, y ∈ s ∧ z ∈ t ∧ y * z = x := mem_image₂ @[simp, norm_cast, to_additive] -lemma coe_mul (s t : finset α) : (↑(s * t) : set α) = ↑s * ↑t := -set.ext $ λ _, by simp only [mem_mul, set.mem_mul, mem_coe] +lemma coe_mul (s t : finset α) : (↑(s * t) : set α) = ↑s * ↑t := coe_image₂ _ _ _ -@[to_additive] -lemma mul_mem_mul (ha : a ∈ s) (hb : b ∈ t) : a * b ∈ s * t := mem_mul.2 ⟨a, b, ha, hb, rfl⟩ - -@[to_additive] -lemma mul_card_le : (s * t).card ≤ s.card * t.card := card_image_le.trans (card_product _ _).le +@[to_additive] lemma mul_mem_mul : a ∈ s → b ∈ t → a * b ∈ s * t := mem_image₂_of_mem +@[to_additive] lemma mul_card_le : (s * t).card ≤ s.card * t.card := card_image₂_le _ _ _ -@[simp, to_additive] lemma empty_mul (s : finset α) : ∅ * s = ∅ := -eq_empty_of_forall_not_mem $ by simp [mem_mul] - -@[simp, to_additive] lemma mul_empty (s : finset α) : s * ∅ = ∅ := -eq_empty_of_forall_not_mem $ by simp [mem_mul] +@[simp, to_additive] lemma empty_mul (s : finset α) : ∅ * s = ∅ := image₂_empty_left +@[simp, to_additive] lemma mul_empty (s : finset α) : s * ∅ = ∅ := image₂_empty_right @[simp, to_additive] lemma mul_nonempty_iff (s t : finset α) : (s * t).nonempty ↔ s.nonempty ∧ t.nonempty := -(nonempty.image_iff _).trans nonempty_product +image₂_nonempty_iff -@[to_additive, mono] lemma mul_subset_mul (hs : s₁ ⊆ s₂) (ht : t₁ ⊆ t₂) : s₁ * t₁ ⊆ s₂ * t₂ := -image_subset_image $ product_subset_product hs ht +@[to_additive] lemma nonempty.mul : s.nonempty → t.nonempty → (s * t).nonempty := nonempty.image₂ -attribute [mono] add_subset_add +@[to_additive, mono] lemma mul_subset_mul : s₁ ⊆ s₂ → t₁ ⊆ t₂ → s₁ * t₁ ⊆ s₂ * t₂ := image₂_subset -@[simp, to_additive] -lemma mul_singleton (a : α) : s * {a} = s.image (* a) := -by { rw [mul_def, product_singleton, map_eq_image, image_image], refl } +attribute [mono] add_subset_add -@[simp, to_additive] -lemma singleton_mul (a : α) : {a} * s = s.image ((*) a) := -by { rw [mul_def, singleton_product, map_eq_image, image_image], refl } +@[simp, to_additive] lemma mul_singleton (a : α) : s * {a} = s.image (* a) := image₂_singleton_right +@[simp, to_additive] lemma singleton_mul (a : α) : {a} * s = s.image ((*) a) := +image₂_singleton_left @[simp, to_additive] -lemma singleton_mul_singleton (a b : α) : ({a} : finset α) * {b} = {a * b} := -by rw [mul_def, singleton_product_singleton, image_singleton] +lemma singleton_mul_singleton (a b : α) : ({a} : finset α) * {b} = {a * b} := image₂_singleton /-- If a finset `u` is contained in the product of two sets `s * t`, we can find two finsets `s'`, `t'` such that `s' ⊆ s`, `t' ⊆ t` and `u ⊆ s' * t'`. -/ @[to_additive "If a finset `u` is contained in the sum of two sets `s + t`, we can find two finsets `s'`, `t'` such that `s' ⊆ s`, `t' ⊆ t` and `u ⊆ s' + t'`."] -lemma subset_mul {s t : set α} (f : ↑u ⊆ s * t) : - ∃ s' t' : finset α, ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ s' * t' := -begin - apply finset.induction_on' u, - { exact ⟨∅, ∅, set.empty_subset _, set.empty_subset _, empty_subset _⟩ }, - rintro a u ha _ _ ⟨s', t', hs, hs', h⟩, - obtain ⟨x, y, hx, hy, ha⟩ := f ha, - use [insert x s', insert y t'], - simp_rw [coe_insert, set.insert_subset], - exact ⟨⟨hx, hs⟩, ⟨hy, hs'⟩, insert_subset.2 ⟨mem_mul.2 ⟨x, y, mem_insert_self _ _, - mem_insert_self _ _, ha⟩, h.trans $ mul_subset_mul (subset_insert _ _) $ subset_insert _ _⟩⟩, -end +lemma subset_mul {s t : set α} : ↑u ⊆ s * t → ∃ s' t' : finset α, ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ s' * t' := +subset_image₂ end has_mul @@ -212,10 +191,10 @@ lemma mul_zero_subset (s : finset α) : s * 0 ⊆ 0 := by simp [subset_iff, mem_ lemma zero_mul_subset (s : finset α) : 0 * s ⊆ 0 := by simp [subset_iff, mem_mul] lemma nonempty.mul_zero (hs : s.nonempty) : s * 0 = 0 := -s.mul_zero_subset.antisymm $ by simpa [finset.mem_mul] using hs +s.mul_zero_subset.antisymm $ by simpa [mem_mul] using hs lemma nonempty.zero_mul (hs : s.nonempty) : 0 * s = 0 := -s.zero_mul_subset.antisymm $ by simpa [finset.mem_mul] using hs +s.zero_mul_subset.antisymm $ by simpa [mem_mul] using hs end mul_zero_class @@ -293,63 +272,40 @@ lemma div_def : s / t = (s.product t).image (λ p : α × α, p.1 / p.2) := rfl @[to_additive add_image_prod] lemma image_div_prod : (s.product t).image (λ x : α × α, x.fst / x.snd) = s / t := rfl -@[to_additive] -lemma mem_div {x : α} : x ∈ s / t ↔ ∃ y z, y ∈ s ∧ z ∈ t ∧ y / z = x := -by simp only [finset.div_def, and.assoc, mem_image, exists_prop, prod.exists, mem_product] +@[to_additive] lemma mem_div : a ∈ s / t ↔ ∃ b c, b ∈ s ∧ c ∈ t ∧ b / c = a := mem_image₂ @[simp, norm_cast, to_additive] -lemma coe_div (s t : finset α) : (↑(s / t) : set α) = ↑s / ↑t := -set.ext $ λ _, by simp only [mem_div, set.mem_div, mem_coe] +lemma coe_div (s t : finset α) : (↑(s / t) : set α) = ↑s / ↑t := coe_image₂ _ _ _ -@[to_additive] -lemma div_mem_div (ha : a ∈ s) (hb : b ∈ t) : a / b ∈ s / t := mem_div.2 ⟨a, b, ha, hb, rfl⟩ - -@[to_additive] -lemma div_card_le : (s / t).card ≤ s.card * t.card := card_image_le.trans (card_product _ _).le +@[to_additive] lemma div_mem_div : a ∈ s → b ∈ t → a / b ∈ s / t := mem_image₂_of_mem +@[to_additive] lemma div_card_le : (s / t).card ≤ s.card * t.card := card_image₂_le _ _ _ -@[simp, to_additive] lemma empty_div (s : finset α) : ∅ / s = ∅ := -eq_empty_of_forall_not_mem $ by simp [mem_div] - -@[simp, to_additive] lemma div_empty (s : finset α) : s / ∅ = ∅ := -eq_empty_of_forall_not_mem $ by simp [mem_div] +@[simp, to_additive] lemma empty_div (s : finset α) : ∅ / s = ∅ := image₂_empty_left +@[simp, to_additive] lemma div_empty (s : finset α) : s / ∅ = ∅ := image₂_empty_right @[simp, to_additive] lemma div_nonempty_iff (s t : finset α) : (s / t).nonempty ↔ s.nonempty ∧ t.nonempty := -(nonempty.image_iff _).trans nonempty_product +image₂_nonempty_iff + +@[to_additive] lemma nonempty.div : s.nonempty → t.nonempty → (s / t).nonempty := nonempty.image₂ -@[to_additive, mono] lemma div_subset_div (hs : s₁ ⊆ s₂) (ht : t₁ ⊆ t₂) : s₁ / t₁ ⊆ s₂ / t₂ := -image_subset_image $ product_subset_product hs ht +@[to_additive, mono] lemma div_subset_div : s₁ ⊆ s₂ → t₁ ⊆ t₂ → s₁ / t₁ ⊆ s₂ / t₂ := image₂_subset attribute [mono] add_subset_add -@[simp, to_additive] -lemma div_singleton (a : α) : s / {a} = s.image (/ a) := -by { rw [div_def, product_singleton, map_eq_image, image_image], refl } +@[simp, to_additive] lemma div_singleton (a : α) : s / {a} = s.image (/ a) := image₂_singleton_right +@[simp, to_additive] lemma singleton_div (a : α) : {a} / s = s.image ((/) a) := +image₂_singleton_left @[simp, to_additive] -lemma singleton_div (a : α) : {a} / s = s.image ((/) a) := -by { rw [div_def, singleton_product, map_eq_image, image_image], refl } - -@[simp, to_additive] -lemma singleton_div_singleton (a b : α) : ({a} : finset α) / {b} = {a / b} := -by rw [div_def, singleton_product_singleton, image_singleton] +lemma singleton_div_singleton (a b : α) : ({a} : finset α) / {b} = {a / b} := image₂_singleton /-- If a finset `u` is contained in the product of two sets `s / t`, we can find two finsets `s'`, `t'` such that `s' ⊆ s`, `t' ⊆ t` and `u ⊆ s' / t'`. -/ @[to_additive "If a finset `u` is contained in the sum of two sets `s - t`, we can find two finsets `s'`, `t'` such that `s' ⊆ s`, `t' ⊆ t` and `u ⊆ s' - t'`."] -lemma subset_div {s t : set α} (f : ↑u ⊆ s / t) : - ∃ s' t' : finset α, ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ s' / t' := -begin - apply finset.induction_on' u, - { exact ⟨∅, ∅, set.empty_subset _, set.empty_subset _, empty_subset _⟩ }, - rintro a u ha _ _ ⟨s', t', hs, hs', h⟩, - obtain ⟨x, y, hx, hy, ha⟩ := f ha, - use [insert x s', insert y t'], - simp_rw [coe_insert, set.insert_subset], - exact ⟨⟨hx, hs⟩, ⟨hy, hs'⟩, insert_subset.2 ⟨mem_div.2 ⟨x, y, mem_insert_self _ _, - mem_insert_self _ _, ha⟩, h.trans $ div_subset_div (subset_insert _ _) $ subset_insert _ _⟩⟩, -end +lemma subset_div {s t : set α} : ↑u ⊆ s / t → ∃ s' t' : finset α, ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ s' / t' := +subset_image₂ end has_div @@ -362,10 +318,10 @@ lemma div_zero_subset (s : finset α) : s / 0 ⊆ 0 := by simp [subset_iff, mem_ lemma zero_div_subset (s : finset α) : 0 / s ⊆ 0 := by simp [subset_iff, mem_div] lemma nonempty.div_zero (hs : s.nonempty) : s / 0 = 0 := -s.div_zero_subset.antisymm $ by simpa [finset.mem_div] using hs +s.div_zero_subset.antisymm $ by simpa [mem_div] using hs lemma nonempty.zero_div (hs : s.nonempty) : 0 / s = 0 := -s.zero_div_subset.antisymm $ by simpa [finset.mem_div] using hs +s.zero_div_subset.antisymm $ by simpa [mem_div] using hs end group_with_zero @@ -472,8 +428,7 @@ variables [decidable_eq β] [has_scalar α β] {s s₁ s₂ : finset α} {t t₁ /-- The pointwise product of two finsets `s` and `t`: `s • t = {x • y | x ∈ s, y ∈ t}`. -/ @[to_additive has_vadd "The pointwise sum of two finsets `s` and `t`: `s +ᵥ t = {x +ᵥ y | x ∈ s, y ∈ t}`."] -protected def has_scalar : has_scalar (finset α) (finset β) := -⟨λ s t, (s.product t).image $ λ p : α × β, p.1 • p.2⟩ +protected def has_scalar : has_scalar (finset α) (finset β) := ⟨image₂ (•)⟩ localized "attribute [instance] finset.has_scalar finset.has_vadd" in pointwise @@ -482,69 +437,44 @@ localized "attribute [instance] finset.has_scalar finset.has_vadd" in pointwise @[to_additive] lemma image_smul_product : (s.product t).image (λ x : α × β, x.fst • x.snd) = s • t := rfl -@[to_additive] -lemma mem_smul {x : β} : x ∈ s • t ↔ ∃ y z, y ∈ s ∧ z ∈ t ∧ y • z = x := -by simp only [finset.smul_def, and.assoc, mem_image, exists_prop, prod.exists, mem_product] +@[to_additive] lemma mem_smul {x : β} : x ∈ s • t ↔ ∃ y z, y ∈ s ∧ z ∈ t ∧ y • z = x := mem_image₂ @[simp, norm_cast, to_additive] lemma coe_smul (s : finset α) (t : finset β) : (↑(s • t) : set β) = (s : set α) • t := -set.ext $ λ _, by simp only [mem_smul, set.mem_smul, mem_coe] +coe_image₂ _ _ _ -@[to_additive] -lemma smul_mem_smul (ha : a ∈ s) (hb : b ∈ t) : a • b ∈ s • t := mem_smul.2 ⟨a, b, ha, hb, rfl⟩ +@[to_additive] lemma smul_mem_smul : a ∈ s → b ∈ t → a • b ∈ s • t := mem_image₂_of_mem +@[to_additive] lemma smul_card_le : (s • t).card ≤ s.card • t.card := card_image₂_le _ _ _ -@[to_additive] -lemma smul_card_le : (s • t).card ≤ s.card • t.card := card_image_le.trans (card_product _ _).le - -@[simp, to_additive] lemma empty_smul (t : finset β) : (∅ : finset α) • t = ∅ := -eq_empty_of_forall_not_mem $ by simp [mem_smul] - -@[simp, to_additive] lemma smul_empty (s : finset α) : s • (∅ : finset β) = ∅ := -eq_empty_of_forall_not_mem $ by simp [mem_smul] +@[simp, to_additive] lemma empty_smul (t : finset β) : (∅ : finset α) • t = ∅ := image₂_empty_left +@[simp, to_additive] lemma smul_empty (s : finset α) : s • (∅ : finset β) = ∅ := image₂_empty_right @[simp, to_additive] -lemma smul_nonempty_iff : (s • t).nonempty ↔ s.nonempty ∧ t.nonempty := -(nonempty.image_iff _).trans nonempty_product +lemma smul_nonempty_iff : (s • t).nonempty ↔ s.nonempty ∧ t.nonempty := image₂_nonempty_iff -@[to_additive] lemma nonempty.smul (hs : s.nonempty) (ht : t.nonempty) : (s • t).nonempty := -smul_nonempty_iff.2 ⟨hs, ht⟩ +@[to_additive] lemma nonempty.smul : s.nonempty → t.nonempty → (s • t).nonempty := nonempty.image₂ -@[to_additive, mono] lemma smul_subset_smul (hs : s₁ ⊆ s₂) (ht : t₁ ⊆ t₂) : s₁ • t₁ ⊆ s₂ • t₂ := -image_subset_image $ product_subset_product hs ht +@[to_additive, mono] lemma smul_subset_smul : s₁ ⊆ s₂ → t₁ ⊆ t₂ → s₁ • t₁ ⊆ s₂ • t₂ := image₂_subset attribute [mono] add_subset_add @[simp, to_additive] -lemma smul_singleton (b : β) : s • ({b} : finset β) = s.image (• b) := -by { classical, rw [smul_def, product_singleton, map_eq_image, image_image], refl } +lemma smul_singleton (b : β) : s • ({b} : finset β) = s.image (• b) := image₂_singleton_right @[simp, to_additive] -lemma singleton_smul (a : α) : ({a} : finset α) • t = t.image ((•) a) := -by { classical, rw [smul_def, singleton_product, map_eq_image, image_image], refl } +lemma singleton_smul (a : α) : ({a} : finset α) • t = t.image ((•) a) := image₂_singleton_left @[simp, to_additive] lemma singleton_smul_singleton (a : α) (b : β) : ({a} : finset α) • ({b} : finset β) = {a • b} := -by rw [smul_def, singleton_product_singleton, image_singleton] +image₂_singleton /-- If a finset `u` is contained in the scalar product of two sets `s • t`, we can find two finsets `s'`, `t'` such that `s' ⊆ s`, `t' ⊆ t` and `u ⊆ s' • t'`. -/ @[to_additive "If a finset `u` is contained in the scalar sum of two sets `s +ᵥ t`, we can find two finsets `s'`, `t'` such that `s' ⊆ s`, `t' ⊆ t` and `u ⊆ s' +ᵥ t'`."] -lemma subset_smul {s : set α} {t : set β} (f : ↑u ⊆ s • t) : - ∃ (s' : finset α) (t' : finset β), ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ s' • t' := -begin - apply finset.induction_on' u, - { exact ⟨∅, ∅, set.empty_subset _, set.empty_subset _, empty_subset _⟩ }, - rintro a u ha _ _ ⟨s', t', hs, hs', h⟩, - obtain ⟨x, y, hx, hy, ha⟩ := f ha, - classical, - use [insert x s', insert y t'], - simp_rw [coe_insert, set.insert_subset], - refine ⟨⟨hx, hs⟩, ⟨hy, hs'⟩, _⟩, - convert insert_subset.2 ⟨mem_smul.2 ⟨x, y, mem_insert_self _ _, - mem_insert_self _ _, ha⟩, h.trans $ _⟩, - convert smul_subset_smul (subset_insert _ _) (subset_insert _ _), -end +lemma subset_smul {s : set α} {t : set β} : + ↑u ⊆ s • t → ∃ (s' : finset α) (t' : finset β), ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ s' • t' := +subset_image₂ end has_scalar @@ -555,69 +485,45 @@ variables [decidable_eq α] [has_vsub α β] {s s₁ s₂ t t₁ t₂ u : finset include α /-- The pointwise product of two finsets `s` and `t`: `s -ᵥ t = {x -ᵥ y | x ∈ s, y ∈ t}`. -/ -protected def has_vsub : has_vsub (finset α) (finset β) := -⟨λ s t, (s.product t).image $ λ p : β × β, p.1 -ᵥ p.2⟩ +protected def has_vsub : has_vsub (finset α) (finset β) := ⟨image₂ (-ᵥ)⟩ localized "attribute [instance] finset.has_vsub" in pointwise -lemma vsub_def : s -ᵥ t = (s.product t).image (λ p : β × β, p.1 -ᵥ p.2) := rfl -lemma image_vsub_product : (s.product t).image (λ x : β × β, x.fst -ᵥ x.snd) = s -ᵥ t := rfl +lemma vsub_def : s -ᵥ t = image₂ (-ᵥ) s t := rfl +@[simp] lemma image_vsub_product : image₂ (-ᵥ) s t = s -ᵥ t := rfl -lemma mem_vsub : a ∈ s -ᵥ t ↔ ∃ b c, b ∈ s ∧ c ∈ t ∧ b -ᵥ c = a := -by simp only [finset.vsub_def, and.assoc, mem_image, exists_prop, prod.exists, mem_product] +lemma mem_vsub : a ∈ s -ᵥ t ↔ ∃ b c, b ∈ s ∧ c ∈ t ∧ b -ᵥ c = a := mem_image₂ @[simp, norm_cast] -lemma coe_vsub (s t : finset β) : (↑(s -ᵥ t) : set α) = (s : set β) -ᵥ t := -set.ext $ λ _, by simp only [mem_vsub, set.mem_vsub, mem_coe] - -lemma vsub_mem_vsub (hb : b ∈ s) (hc : c ∈ t) : b -ᵥ c ∈ s -ᵥ t := mem_vsub.2 ⟨b, c, hb, hc, rfl⟩ +lemma coe_vsub (s t : finset β) : (↑(s -ᵥ t) : set α) = (s : set β) -ᵥ t := coe_image₂ _ _ _ -lemma vsub_card_le : (s -ᵥ t : finset α).card ≤ s.card * t.card := -card_image_le.trans (card_product _ _).le +lemma vsub_mem_vsub : b ∈ s → c ∈ t → b -ᵥ c ∈ s -ᵥ t := mem_image₂_of_mem +lemma vsub_card_le : (s -ᵥ t : finset α).card ≤ s.card * t.card := card_image₂_le _ _ _ -@[simp] lemma empty_vsub (t : finset β) : (∅ : finset β) -ᵥ t = ∅ := -eq_empty_of_forall_not_mem $ by simp [mem_vsub] +@[simp] lemma empty_vsub (t : finset β) : (∅ : finset β) -ᵥ t = ∅ := image₂_empty_left +@[simp] lemma vsub_empty (s : finset β) : s -ᵥ (∅ : finset β) = ∅ := image₂_empty_right -@[simp] lemma vsub_empty (s : finset β) : s -ᵥ (∅ : finset β) = ∅ := -eq_empty_of_forall_not_mem $ by simp [mem_vsub] - -@[simp] -lemma vsub_nonempty_iff : (s -ᵥ t : finset α).nonempty ↔ s.nonempty ∧ t.nonempty := -(nonempty.image_iff _).trans nonempty_product +@[simp] lemma vsub_nonempty_iff : (s -ᵥ t : finset α).nonempty ↔ s.nonempty ∧ t.nonempty := +image₂_nonempty_iff -lemma nonempty.vsub (hs : s.nonempty) (ht : t.nonempty) : (s -ᵥ t : finset α).nonempty := -vsub_nonempty_iff.2 ⟨hs, ht⟩ +lemma nonempty.vsub : s.nonempty → t.nonempty → (s -ᵥ t : finset α).nonempty := nonempty.image₂ -@[mono] lemma vsub_subset_vsub (hs : s₁ ⊆ s₂) (ht : t₁ ⊆ t₂) : s₁ -ᵥ t₁ ⊆ s₂ -ᵥ t₂ := -image_subset_image $ product_subset_product hs ht +@[mono] lemma vsub_subset_vsub : s₁ ⊆ s₂ → t₁ ⊆ t₂ → s₁ -ᵥ t₁ ⊆ s₂ -ᵥ t₂ := image₂_subset @[simp] lemma vsub_singleton (b : β) : s -ᵥ ({b} : finset β) = s.image (-ᵥ b) := -by { classical, rw [vsub_def, product_singleton, map_eq_image, image_image], refl } +image₂_singleton_right @[simp] lemma singleton_vsub (a : β) : ({a} : finset β) -ᵥ t = t.image ((-ᵥ) a) := -by { classical, rw [vsub_def, singleton_product, map_eq_image, image_image], refl } +image₂_singleton_left @[simp] -lemma singleton_vsub_singleton (a b : β) : ({a} : finset β) -ᵥ {b} = {a -ᵥ b} := -by rw [vsub_def, singleton_product_singleton, image_singleton] +lemma singleton_vsub_singleton (a b : β) : ({a} : finset β) -ᵥ {b} = {a -ᵥ b} := image₂_singleton /-- If a finset `u` is contained in the pointwise subtraction of two sets `s -ᵥ t`, we can find two finsets `s'`, `t'` such that `s' ⊆ s`, `t' ⊆ t` and `u ⊆ s' -ᵥ t'`. -/ -lemma subset_vsub {s t : set β} {u : finset α} (f : ↑u ⊆ s -ᵥ t) : - ∃ s' t' : finset β, ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ s' -ᵥ t' := -begin - apply finset.induction_on' u, - { exact ⟨∅, ∅, set.empty_subset _, set.empty_subset _, empty_subset _⟩ }, - rintro a u ha _ _ ⟨s', t', hs, hs', h⟩, - obtain ⟨x, y, hx, hy, ha⟩ := f ha, - classical, - use [insert x s', insert y t'], - simp_rw [coe_insert, set.insert_subset], - refine ⟨⟨hx, hs⟩, ⟨hy, hs'⟩, _⟩, - convert insert_subset.2 ⟨mem_vsub.2 ⟨x, y, mem_insert_self _ _, - mem_insert_self _ _, ha⟩, h.trans $ _⟩, - convert vsub_subset_vsub (subset_insert _ _) (subset_insert _ _), -end +lemma subset_vsub {s t : set β} {u : finset α} : + ↑u ⊆ s -ᵥ t → ∃ s' t' : finset β, ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ s' -ᵥ t' := +subset_image₂ end has_vsub @@ -645,9 +551,7 @@ by simp only [finset.smul_finset_def, and.assoc, mem_image, exists_prop, prod.ex @[simp, norm_cast, to_additive] lemma coe_smul_finset (s : finset β) : (↑(a • s) : set β) = a • s := coe_image -@[to_additive] -lemma smul_finset_mem_smul_finset (hb : b ∈ s) : a • b ∈ a • s := mem_image_of_mem _ hb - +@[to_additive] lemma smul_finset_mem_smul_finset : b ∈ s → a • b ∈ a • s := mem_image_of_mem _ @[to_additive] lemma smul_finset_card_le : (a • s).card ≤ s.card := card_image_le @[simp, to_additive] lemma smul_finset_empty (a : α) : a • (∅ : finset β) = ∅ := image_empty _ diff --git a/src/order/filter/n_ary.lean b/src/order/filter/n_ary.lean index 70906967e9f1c..7feb6c067c9a0 100644 --- a/src/order/filter/n_ary.lean +++ b/src/order/filter/n_ary.lean @@ -18,7 +18,8 @@ operations on filters. ## Notes -This file is very similar to the n-ary section of `data.set.basic`. Please keep them in sync. +This file is very similar to the n-ary section of `data.set.basic` and to `data.finset.n_ary`. +Please keep them in sync. -/ open function set From 1c39267116733f40fe874efaf25d8571def0d0fc Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 3 May 2022 11:03:24 +0000 Subject: [PATCH 030/153] ci(elan): update dead repository URLs (#13906) `Kha/elan` is redirected by github to `leanprover/elan`, but seemingly with a cache that is delayed. [Zulip](https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/install.20elan.20fails.20in.20CI/near/280981154) --- .github/workflows/bors.yml | 4 ++-- .github/workflows/build.yml | 4 ++-- .github/workflows/build.yml.in | 4 ++-- .github/workflows/build_fork.yml | 4 ++-- 4 files changed, 8 insertions(+), 8 deletions(-) diff --git a/.github/workflows/bors.yml b/.github/workflows/bors.yml index 7a8e24d447787..5a2e9fed9d73d 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -142,7 +142,7 @@ jobs: - name: install elan run: | set -o pipefail - curl https://raw.githubusercontent.com/Kha/elan/master/elan-init.sh -sSf | sh -s -- --default-toolchain none -y + curl https://raw.githubusercontent.com/leanprover/elan/master/elan-init.sh -sSf | sh -s -- --default-toolchain none -y ~/.elan/bin/lean --version echo "$HOME/.elan/bin" >> $GITHUB_PATH @@ -179,7 +179,7 @@ jobs: - name: install elan run: | set -o pipefail - curl https://raw.githubusercontent.com/Kha/elan/master/elan-init.sh -sSf | sh -s -- --default-toolchain none -y + curl https://raw.githubusercontent.com/leanprover/elan/master/elan-init.sh -sSf | sh -s -- --default-toolchain none -y ~/.elan/bin/lean --version echo "$HOME/.elan/bin" >> $GITHUB_PATH diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index a43a55cd629c3..7914e3a217905 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -150,7 +150,7 @@ jobs: - name: install elan run: | set -o pipefail - curl https://raw.githubusercontent.com/Kha/elan/master/elan-init.sh -sSf | sh -s -- --default-toolchain none -y + curl https://raw.githubusercontent.com/leanprover/elan/master/elan-init.sh -sSf | sh -s -- --default-toolchain none -y ~/.elan/bin/lean --version echo "$HOME/.elan/bin" >> $GITHUB_PATH @@ -187,7 +187,7 @@ jobs: - name: install elan run: | set -o pipefail - curl https://raw.githubusercontent.com/Kha/elan/master/elan-init.sh -sSf | sh -s -- --default-toolchain none -y + curl https://raw.githubusercontent.com/leanprover/elan/master/elan-init.sh -sSf | sh -s -- --default-toolchain none -y ~/.elan/bin/lean --version echo "$HOME/.elan/bin" >> $GITHUB_PATH diff --git a/.github/workflows/build.yml.in b/.github/workflows/build.yml.in index 36f4763d5f939..4a28319cc146a 100644 --- a/.github/workflows/build.yml.in +++ b/.github/workflows/build.yml.in @@ -128,7 +128,7 @@ jobs: - name: install elan run: | set -o pipefail - curl https://raw.githubusercontent.com/Kha/elan/master/elan-init.sh -sSf | sh -s -- --default-toolchain none -y + curl https://raw.githubusercontent.com/leanprover/elan/master/elan-init.sh -sSf | sh -s -- --default-toolchain none -y ~/.elan/bin/lean --version echo "$HOME/.elan/bin" >> $GITHUB_PATH @@ -165,7 +165,7 @@ jobs: - name: install elan run: | set -o pipefail - curl https://raw.githubusercontent.com/Kha/elan/master/elan-init.sh -sSf | sh -s -- --default-toolchain none -y + curl https://raw.githubusercontent.com/leanprover/elan/master/elan-init.sh -sSf | sh -s -- --default-toolchain none -y ~/.elan/bin/lean --version echo "$HOME/.elan/bin" >> $GITHUB_PATH diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index d601128dee83a..07e788387667d 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -148,7 +148,7 @@ jobs: - name: install elan run: | set -o pipefail - curl https://raw.githubusercontent.com/Kha/elan/master/elan-init.sh -sSf | sh -s -- --default-toolchain none -y + curl https://raw.githubusercontent.com/leanprover/elan/master/elan-init.sh -sSf | sh -s -- --default-toolchain none -y ~/.elan/bin/lean --version echo "$HOME/.elan/bin" >> $GITHUB_PATH @@ -185,7 +185,7 @@ jobs: - name: install elan run: | set -o pipefail - curl https://raw.githubusercontent.com/Kha/elan/master/elan-init.sh -sSf | sh -s -- --default-toolchain none -y + curl https://raw.githubusercontent.com/leanprover/elan/master/elan-init.sh -sSf | sh -s -- --default-toolchain none -y ~/.elan/bin/lean --version echo "$HOME/.elan/bin" >> $GITHUB_PATH From 65cad4187da45259d3d74f3f2de5139ee0e8df8d Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Tue, 3 May 2022 11:43:28 +0000 Subject: [PATCH 031/153] chore(.github/workflows): use separate secret token for dependent issues (#13902) Co-authored-by: Scott Morrison --- .github/workflows/dependent-issues.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.github/workflows/dependent-issues.yml b/.github/workflows/dependent-issues.yml index ac19ed5287e83..48e6eace562e9 100644 --- a/.github/workflows/dependent-issues.yml +++ b/.github/workflows/dependent-issues.yml @@ -32,7 +32,7 @@ jobs: - uses: z0al/dependent-issues@v1 env: # (Required) The token to use to make API calls to GitHub. - GITHUB_TOKEN: ${{ secrets.TRIAGE_TOKEN }} + GITHUB_TOKEN: ${{ secrets.DEPENDENT_ISSUES_TOKEN }} with: # (Optional) The label to use to mark dependent issues label: blocked-by-other-PR From 7931ba4416ad13f44768e3aeb95f834e546f5cbf Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 3 May 2022 12:18:49 +0000 Subject: [PATCH 032/153] feat(order/conditionally_complete_lattice): Simp theorems (#13756) We remove `supr_unit` and `infi_unit` since, thanks to #13741, they can be proven by `simp`. --- src/order/conditionally_complete_lattice.lean | 10 ++-------- 1 file changed, 2 insertions(+), 8 deletions(-) diff --git a/src/order/conditionally_complete_lattice.lean b/src/order/conditionally_complete_lattice.lean index ffa8d0dc2f6b9..cbf266460a395 100644 --- a/src/order/conditionally_complete_lattice.lean +++ b/src/order/conditionally_complete_lattice.lean @@ -501,19 +501,13 @@ by rw [supr, range_const, cSup_singleton] @[simp] theorem cinfi_const [hι : nonempty ι] {a : α} : (⨅ b:ι, a) = a := @csupr_const (order_dual α) _ _ _ _ -theorem supr_unique [unique ι] {s : ι → α} : (⨆ i, s i) = s default := +@[simp] theorem supr_unique [unique ι] {s : ι → α} : (⨆ i, s i) = s default := have ∀ i, s i = s default := λ i, congr_arg s (unique.eq_default i), by simp only [this, csupr_const] -theorem infi_unique [unique ι] {s : ι → α} : (⨅ i, s i) = s default := +@[simp] theorem infi_unique [unique ι] {s : ι → α} : (⨅ i, s i) = s default := @supr_unique (order_dual α) _ _ _ _ -@[simp] theorem supr_unit {f : unit → α} : (⨆ x, f x) = f () := -by { convert supr_unique, apply_instance } - -@[simp] theorem infi_unit {f : unit → α} : (⨅ x, f x) = f () := -@supr_unit (order_dual α) _ _ - @[simp] lemma csupr_pos {p : Prop} {f : p → α} (hp : p) : (⨆ h : p, f h) = f hp := by haveI := unique_prop hp; exact supr_unique From 72816f95e620cc25e6f55468cf6eca1567e30041 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 3 May 2022 12:18:50 +0000 Subject: [PATCH 033/153] feat(data/real/nnreal): add `nnreal.forall` and `nnreal.exists` (#13774) --- src/data/real/nnreal.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/src/data/real/nnreal.lean b/src/data/real/nnreal.lean index b9e8f7b77ceaa..2c66a26ed9684 100644 --- a/src/data/real/nnreal.lean +++ b/src/data/real/nnreal.lean @@ -80,6 +80,12 @@ iff.intro nnreal.eq (congr_arg coe) lemma ne_iff {x y : ℝ≥0} : (x : ℝ) ≠ (y : ℝ) ↔ x ≠ y := not_iff_not_of_iff $ nnreal.eq_iff +protected lemma «forall» {p : ℝ≥0 → Prop} : (∀ x : ℝ≥0, p x) ↔ ∀ (x : ℝ) (hx : 0 ≤ x), p ⟨x, hx⟩ := +subtype.forall + +protected lemma «exists» {p : ℝ≥0 → Prop} : (∃ x : ℝ≥0, p x) ↔ ∃ (x : ℝ) (hx : 0 ≤ x), p ⟨x, hx⟩ := +subtype.exists + /-- Reinterpret a real number `r` as a non-negative real number. Returns `0` if `r < 0`. -/ noncomputable def _root_.real.to_nnreal (r : ℝ) : ℝ≥0 := ⟨max r 0, le_max_right _ _⟩ From b07c0f70877ebe529e60f0cf5f0f3ecad4f3dbdd Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 3 May 2022 12:18:51 +0000 Subject: [PATCH 034/153] feat(set_theory/game/basic): Add `le_rfl` on games (#13814) --- src/set_theory/game/basic.lean | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/src/set_theory/game/basic.lean b/src/set_theory/game/basic.lean index ae163628b2e02..2626b6765f8c8 100644 --- a/src/set_theory/game/basic.lean +++ b/src/set_theory/game/basic.lean @@ -57,8 +57,10 @@ instance : has_le game := -- Adding `@[refl]` and `@[trans]` attributes here would override the ones on -- `preorder.le_refl` and `preorder.le_trans`, which breaks all non-`game` uses of `≤`! -theorem le_refl : ∀ x : game, x ≤ x := -by { rintro ⟨x⟩, apply pgame.le_refl } +theorem le_rfl : ∀ {x : game}, x ≤ x := +by { rintro ⟨x⟩, exact pgame.le_rfl } +theorem le_refl (x : game) : x ≤ x := +le_rfl theorem le_trans : ∀ x y z : game, x ≤ y → y ≤ z → x ≤ z := by { rintro ⟨x⟩ ⟨y⟩ ⟨z⟩, apply pgame.le_trans } theorem le_antisymm : ∀ x y : game, x ≤ y → y ≤ x → x = y := From f6cb9be24263235f5326e2fc1fb4fea8df7c306e Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 3 May 2022 12:18:52 +0000 Subject: [PATCH 035/153] fix(data/complex/basic): make complex addition computable again (#13837) This was fixed once before in #8166 (5f2358c43b769b334f3986a96565e606fe5bccec), but a new noncomputable shortcut appears if your file has more imports. --- src/data/complex/basic.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/src/data/complex/basic.lean b/src/data/complex/basic.lean index 7baeb45472365..a68b5c91b8155 100644 --- a/src/data/complex/basic.lean +++ b/src/data/complex/basic.lean @@ -173,6 +173,10 @@ by refine_struct zsmul := λ n z, ⟨n • z.re - 0 * z.im, n • z.im + 0 * z.re⟩ }; intros; try { refl }; apply ext_iff.2; split; simp; {ring1 <|> ring_nf} +/-- This shortcut instance ensures we do not find `add_comm_group` via the noncomputable +`complex.normed_group` instance. -/ +instance : add_comm_group ℂ := by apply_instance + /-- This shortcut instance ensures we do not find `ring` via the noncomputable `complex.field` instance. -/ instance : ring ℂ := by apply_instance From e104992d2ed6f4c774f8f35dd84855077711e7cf Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 3 May 2022 12:18:53 +0000 Subject: [PATCH 036/153] chore(order/*): Replace total partial orders by linear orders (#13839) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit `partial_order α` + `is_total α (≤)` has no more theorems than `linear_order α` but is nonetheless used in some places. This replaces those uses by `linear_order α` or `complete_linear_order α`. Also make implicit one argument of `finset.lt_inf'_iff` and friends. --- src/analysis/locally_convex/weak_dual.lean | 7 +- src/data/finset/lattice.lean | 186 +++++++++++--------- src/data/real/nnreal.lean | 2 +- src/data/set/intervals/basic.lean | 68 ++++--- src/order/complete_lattice.lean | 3 + src/order/lattice.lean | 95 +++++----- src/order/min_max.lean | 36 ++-- src/order/omega_complete_partial_order.lean | 40 ++--- src/topology/algebra/order/basic.lean | 20 +-- 9 files changed, 232 insertions(+), 225 deletions(-) diff --git a/src/analysis/locally_convex/weak_dual.lean b/src/analysis/locally_convex/weak_dual.lean index 903938767bf15..164e1d336a29c 100644 --- a/src/analysis/locally_convex/weak_dual.lean +++ b/src/analysis/locally_convex/weak_dual.lean @@ -101,11 +101,8 @@ begin let U' := hU₁.to_finset, by_cases hU₃ : U.fst.nonempty, { have hU₃' : U'.nonempty := (set.finite.to_finset.nonempty hU₁).mpr hU₃, - let r := U'.inf' hU₃' U.snd, - have hr : 0 < r := - (finset.lt_inf'_iff hU₃' _).mpr (λ y hy, hU₂ y ((set.finite.mem_to_finset hU₁).mp hy)), - use [seminorm.ball (U'.sup p) (0 : E) r], - refine ⟨p.basis_sets_mem _ hr, λ x hx y hy, _⟩, + refine ⟨(U'.sup p).ball 0 $ U'.inf' hU₃' U.snd, p.basis_sets_mem _ $ + (finset.lt_inf'_iff _).2 $ λ y hy, hU₂ y $ (hU₁.mem_to_finset).mp hy, λ x hx y hy, _⟩, simp only [set.mem_preimage, set.mem_pi, mem_ball_zero_iff], rw seminorm.mem_ball_zero at hx, rw ←linear_map.to_seminorm_family_apply, diff --git a/src/data/finset/lattice.lean b/src/data/finset/lattice.lean index fc665e5e9deed..5880a4729d6f1 100644 --- a/src/data/finset/lattice.lean +++ b/src/data/finset/lattice.lean @@ -103,22 +103,6 @@ sup_le (λ b hb, le_trans (h b hb) (le_sup hb)) lemma sup_mono (h : s₁ ⊆ s₂) : s₁.sup f ≤ s₂.sup f := sup_le $ assume b hb, le_sup (h hb) -@[simp] lemma sup_lt_iff [is_total α (≤)] {a : α} (ha : ⊥ < a) : s.sup f < a ↔ (∀ b ∈ s, f b < a) := -⟨(λ hs b hb, lt_of_le_of_lt (le_sup hb) hs), finset.cons_induction_on s (λ _, ha) - (λ c t hc, by simpa only [sup_cons, sup_lt_iff, mem_cons, forall_eq_or_imp] using and.imp_right)⟩ - -@[simp] lemma le_sup_iff [is_total α (≤)] {a : α} (ha : ⊥ < a) : a ≤ s.sup f ↔ ∃ b ∈ s, a ≤ f b := -⟨finset.cons_induction_on s (λ h, absurd h (not_le_of_lt ha)) - (λ c t hc ih, by simpa using @or.rec _ _ (∃ b, (b = c ∨ b ∈ t) ∧ a ≤ f b) - (λ h, ⟨c, or.inl rfl, h⟩) (λ h, let ⟨b, hb, hle⟩ := ih h in ⟨b, or.inr hb, hle⟩)), -(λ ⟨b, hb, hle⟩, trans hle (le_sup hb))⟩ - -@[simp] lemma lt_sup_iff [is_total α (≤)] {a : α} : a < s.sup f ↔ ∃ b ∈ s, a < f b := -⟨finset.cons_induction_on s (λ h, absurd h not_lt_bot) - (λ c t hc ih, by simpa using @or.rec _ _ (∃ b, (b = c ∨ b ∈ t) ∧ a < f b) - (λ h, ⟨c, or.inl rfl, h⟩) (λ h, let ⟨b, hb, hlt⟩ := ih h in ⟨b, or.inr hb, hlt⟩)), -(λ ⟨b, hb, hlt⟩, lt_of_lt_of_le hlt (le_sup hb))⟩ - lemma sup_comm (s : finset β) (t : finset γ) (f : β → γ → α) : s.sup (λ b, t.sup (f b)) = t.sup (λ c, s.sup (λ b, f b c)) := begin @@ -168,10 +152,6 @@ lemma comp_sup_eq_sup_comp [semilattice_sup γ] [order_bot γ] {s : finset β} g (s.sup f) = s.sup (g ∘ f) := finset.cons_induction_on s bot (λ c t hc ih, by rw [sup_cons, sup_cons, g_sup, ih]) -lemma comp_sup_eq_sup_comp_of_is_total [is_total α (≤)] {γ : Type} [semilattice_sup γ] [order_bot γ] - (g : α → γ) (mono_g : monotone g) (bot : g ⊥ = ⊥) : g (s.sup f) = s.sup (g ∘ f) := -comp_sup_eq_sup_comp g mono_g.map_sup bot - /-- Computing `sup` in a subtype (closed under `sup`) is the same as computing it in `α`. -/ lemma sup_coe {P : α → Prop} {Pbot : P ⊥} {Psup : ∀{{x y}}, P x → P y → P (x ⊔ y)} @@ -326,15 +306,6 @@ le_inf (λ b hb, le_trans (inf_le hb) (h b hb)) lemma inf_mono (h : s₁ ⊆ s₂) : s₂.inf f ≤ s₁.inf f := le_inf $ assume b hb, inf_le (h hb) -@[simp] lemma lt_inf_iff [is_total α (≤)] {a : α} (ha : a < ⊤) : a < s.inf f ↔ (∀ b ∈ s, a < f b) := -@sup_lt_iff (order_dual α) _ _ _ _ _ _ _ ha - -@[simp] lemma inf_le_iff [is_total α (≤)] {a : α} (ha : a < ⊤) : s.inf f ≤ a ↔ (∃ b ∈ s, f b ≤ a) := -@le_sup_iff (order_dual α) _ _ _ _ _ _ _ ha - -@[simp] lemma inf_lt_iff [is_total α (≤)] {a : α} : s.inf f < a ↔ (∃ b ∈ s, f b < a) := -@lt_sup_iff (order_dual α) _ _ _ _ _ _ _ - lemma inf_attach (s : finset β) (f : β → α) : s.attach.inf (λ x, f x) = s.inf f := @sup_attach (order_dual α) _ _ _ _ _ @@ -384,10 +355,6 @@ lemma comp_inf_eq_inf_comp [semilattice_inf γ] [order_top γ] {s : finset β} g (s.inf f) = s.inf (g ∘ f) := @comp_sup_eq_sup_comp (order_dual α) _ (order_dual γ) _ _ _ _ _ _ _ g_inf top -lemma comp_inf_eq_inf_comp_of_is_total [h : is_total α (≤)] {γ : Type} [semilattice_inf γ] - [order_top γ] (g : α → γ) (mono_g : monotone g) (top : g ⊤ = ⊤) : g (s.inf f) = s.inf (g ∘ f) := -comp_inf_eq_inf_comp g mono_g.map_inf top - /-- Computing `inf` in a subtype (closed under `inf`) is the same as computing it in `α`. -/ lemma inf_coe {P : α → Prop} {Ptop : P ⊤} {Pinf : ∀{{x y}}, P x → P y → P (x ⊓ y)} @@ -452,6 +419,53 @@ lemma inf_sup_distrib_right (s : finset ι) (f : ι → α) (a : α) : end order_top end distrib_lattice +section linear_order +variables [linear_order α] + +section order_bot +variables [order_bot α] {s : finset ι} {f : ι → α} {a : α} + +lemma comp_sup_eq_sup_comp_of_is_total [semilattice_sup β] [order_bot β] (g : α → β) + (mono_g : monotone g) (bot : g ⊥ = ⊥) : g (s.sup f) = s.sup (g ∘ f) := +comp_sup_eq_sup_comp g mono_g.map_sup bot + +@[simp] protected lemma le_sup_iff (ha : ⊥ < a) : a ≤ s.sup f ↔ ∃ b ∈ s, a ≤ f b := +⟨finset.cons_induction_on s (λ h, absurd h (not_le_of_lt ha)) + (λ c t hc ih, by simpa using @or.rec _ _ (∃ b, (b = c ∨ b ∈ t) ∧ a ≤ f b) + (λ h, ⟨c, or.inl rfl, h⟩) (λ h, let ⟨b, hb, hle⟩ := ih h in ⟨b, or.inr hb, hle⟩)), +(λ ⟨b, hb, hle⟩, trans hle (le_sup hb))⟩ + +@[simp] protected lemma lt_sup_iff : a < s.sup f ↔ ∃ b ∈ s, a < f b := +⟨finset.cons_induction_on s (λ h, absurd h not_lt_bot) + (λ c t hc ih, by simpa using @or.rec _ _ (∃ b, (b = c ∨ b ∈ t) ∧ a < f b) + (λ h, ⟨c, or.inl rfl, h⟩) (λ h, let ⟨b, hb, hlt⟩ := ih h in ⟨b, or.inr hb, hlt⟩)), +(λ ⟨b, hb, hlt⟩, lt_of_lt_of_le hlt (le_sup hb))⟩ + +@[simp] protected lemma sup_lt_iff (ha : ⊥ < a) : s.sup f < a ↔ ∀ b ∈ s, f b < a := +⟨(λ hs b hb, lt_of_le_of_lt (le_sup hb) hs), finset.cons_induction_on s (λ _, ha) + (λ c t hc, by simpa only [sup_cons, sup_lt_iff, mem_cons, forall_eq_or_imp] using and.imp_right)⟩ + +end order_bot + +section order_top +variables [order_top α] {s : finset ι} {f : ι → α} {a : α} + +lemma comp_inf_eq_inf_comp_of_is_total [semilattice_inf β] [order_top β] (g : α → β) + (mono_g : monotone g) (top : g ⊤ = ⊤) : g (s.inf f) = s.inf (g ∘ f) := +comp_inf_eq_inf_comp g mono_g.map_inf top + +@[simp] protected lemma inf_le_iff (ha : a < ⊤) : s.inf f ≤ a ↔ ∃ b ∈ s, f b ≤ a := +@finset.le_sup_iff (order_dual α) _ _ _ _ _ _ ha + +@[simp] protected lemma inf_lt_iff : s.inf f < a ↔ (∃ b ∈ s, f b < a) := +@finset.lt_sup_iff (order_dual α) _ _ _ _ _ _ + +@[simp] protected lemma lt_inf_iff (ha : a < ⊤) : a < s.inf f ↔ ∀ b ∈ s, a < f b := +@finset.sup_lt_iff (order_dual α) _ _ _ _ _ _ ha + +end order_top +end linear_order + lemma inf_eq_infi [complete_lattice β] (s : finset α) (f : α → β) : s.inf f = (⨅a∈s, f a) := @sup_eq_supr _ (order_dual β) _ _ _ @@ -512,24 +526,6 @@ end @[simp] lemma sup'_le_iff {a : α} : s.sup' H f ≤ a ↔ ∀ b ∈ s, f b ≤ a := iff.intro (λ h b hb, trans (le_sup' f hb) h) (sup'_le H f) -@[simp] lemma sup'_lt_iff [is_total α (≤)] {a : α} : s.sup' H f < a ↔ (∀ b ∈ s, f b < a) := -begin - rw [←with_bot.coe_lt_coe, coe_sup', sup_lt_iff (with_bot.bot_lt_coe a)], - exact ball_congr (λ b hb, with_bot.coe_lt_coe), -end - -@[simp] lemma le_sup'_iff [is_total α (≤)] {a : α} : a ≤ s.sup' H f ↔ (∃ b ∈ s, a ≤ f b) := -begin - rw [←with_bot.coe_le_coe, coe_sup', le_sup_iff (with_bot.bot_lt_coe a)], - exact bex_congr (λ b hb, with_bot.coe_le_coe), -end - -@[simp] lemma lt_sup'_iff [is_total α (≤)] {a : α} : a < s.sup' H f ↔ (∃ b ∈ s, a < f b) := -begin - rw [←with_bot.coe_lt_coe, coe_sup', lt_sup_iff], - exact bex_congr (λ b hb, with_bot.coe_lt_coe), -end - lemma sup'_bUnion [decidable_eq β] {s : finset γ} (Hs : s.nonempty) {t : γ → finset β} (Ht : ∀ b, (t b).nonempty) : (s.bUnion t).sup' (Hs.bUnion (λ b _, Ht b)) f = s.sup' Hs (λ b, (t b).sup' (Ht b) f) := @@ -563,17 +559,6 @@ begin exacts [h₁, hp a₁ h₁ a₂ h₂] end -lemma exists_mem_eq_sup' [is_total α (≤)] : ∃ b, b ∈ s ∧ s.sup' H f = f b := -begin - refine H.cons_induction (λ c, _) (λ c s hc hs ih, _), - { exact ⟨c, mem_singleton_self c, rfl⟩, }, - { rcases ih with ⟨b, hb, h'⟩, - rw [sup'_cons hs, h'], - cases total_of (≤) (f b) (f c) with h h, - { exact ⟨c, mem_cons.2 (or.inl rfl), sup_eq_left.2 h⟩, }, - { exact ⟨b, mem_cons.2 (or.inr hb), sup_eq_right.2 h⟩, }, }, -end - lemma sup'_mem (s : set α) (w : ∀ x y ∈ s, x ⊔ y ∈ s) {ι : Type*} (t : finset ι) (H : t.nonempty) (p : ι → α) (h : ∀ i ∈ t, p i ∈ s) : @@ -627,18 +612,6 @@ lemma inf'_le {b : β} (h : b ∈ s) : s.inf' ⟨b, h⟩ f ≤ f b := @[simp] lemma inf'_const (a : α) : s.inf' H (λ b, a) = a := @sup'_const (order_dual α) _ _ _ _ _ -@[simp] lemma le_inf'_iff {a : α} : a ≤ s.inf' H f ↔ ∀ b ∈ s, a ≤ f b := -@sup'_le_iff (order_dual α) _ _ _ H f _ - -@[simp] lemma lt_inf'_iff [is_total α (≤)] {a : α} : a < s.inf' H f ↔ (∀ b ∈ s, a < f b) := -@sup'_lt_iff (order_dual α) _ _ _ H f _ _ - -@[simp] lemma inf'_le_iff [is_total α (≤)] {a : α} : s.inf' H f ≤ a ↔ (∃ b ∈ s, f b ≤ a) := -@le_sup'_iff (order_dual α) _ _ _ H f _ _ - -@[simp] lemma inf'_lt_iff [is_total α (≤)] {a : α} : s.inf' H f < a ↔ (∃ b ∈ s, f b < a) := -@lt_sup'_iff (order_dual α) _ _ _ 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) := @@ -653,9 +626,6 @@ lemma inf'_induction {p : α → Prop} (hp : ∀ a₁, p a₁ → ∀ a₂, p a (hs : ∀ b ∈ s, p (f b)) : p (s.inf' H f) := @sup'_induction (order_dual α) _ _ _ H f _ hp hs -lemma exists_mem_eq_inf' [is_total α (≤)] : ∃ b, b ∈ s ∧ s.inf' H f = f b := -@exists_mem_eq_sup' (order_dual α) _ _ _ H f _ - lemma inf'_mem (s : set α) (w : ∀ x y ∈ s, x ⊓ y ∈ s) {ι : Type*} (t : finset ι) (H : t.nonempty) (p : ι → α) (h : ∀ i ∈ t, p i ∈ s) : t.inf' H p ∈ s := inf'_induction H p w h @@ -676,10 +646,6 @@ lemma sup_closed_of_sup_closed {s : set α} (t : finset α) (htne : t.nonempty) (h : ∀ a b ∈ s, a ⊔ b ∈ s) : t.sup id ∈ s := sup'_eq_sup htne id ▸ sup'_induction _ _ h h_subset -lemma exists_mem_eq_sup [is_total α (≤)] (s : finset β) (h : s.nonempty) (f : β → α) : - ∃ b, b ∈ s ∧ s.sup f = f b := -sup'_eq_sup h f ▸ exists_mem_eq_sup' h f - lemma coe_sup_of_nonempty {s : finset β} (h : s.nonempty) (f : β → α) : (↑(s.sup f) : with_bot α) = s.sup (coe ∘ f) := by simp only [←sup'_eq_sup h, coe_sup' h] @@ -696,10 +662,6 @@ lemma inf_closed_of_inf_closed {s : set α} (t : finset α) (htne : t.nonempty) (h : ∀ a b ∈ s, a ⊓ b ∈ s) : t.inf id ∈ s := @sup_closed_of_sup_closed (order_dual α) _ _ _ t htne h_subset h -lemma exists_mem_eq_inf [is_total α (≤)] (s : finset β) (h : s.nonempty) (f : β → α) : - ∃ a, a ∈ s ∧ s.inf f = f a := -@exists_mem_eq_sup (order_dual α) _ _ _ _ _ h f - lemma coe_inf_of_nonempty {s : finset β} (h : s.nonempty) (f : β → α): (↑(s.inf f) : with_top α) = s.inf (λ i, f i) := @coe_sup_of_nonempty (order_dual α) _ _ _ _ h f @@ -746,6 +708,60 @@ protected lemma inf'_apply {s : finset α} (H : s.nonempty) (f : α → Π (b : end inf' +section linear_order +variables [linear_order α] {s : finset ι} (H : s.nonempty) {f : ι → α} {a : α} + +@[simp] lemma le_sup'_iff : a ≤ s.sup' H f ↔ ∃ b ∈ s, a ≤ f b := +begin + rw [←with_bot.coe_le_coe, coe_sup', finset.le_sup_iff (with_bot.bot_lt_coe a)], + exact bex_congr (λ b hb, with_bot.coe_le_coe), +end + +@[simp] lemma lt_sup'_iff : a < s.sup' H f ↔ ∃ b ∈ s, a < f b := +begin + rw [←with_bot.coe_lt_coe, coe_sup', finset.lt_sup_iff], + exact bex_congr (λ b hb, with_bot.coe_lt_coe), +end + +@[simp] lemma sup'_lt_iff : s.sup' H f < a ↔ ∀ i ∈ s, f i < a := +begin + rw [←with_bot.coe_lt_coe, coe_sup', finset.sup_lt_iff (with_bot.bot_lt_coe a)], + exact ball_congr (λ b hb, with_bot.coe_lt_coe), +end + +@[simp] lemma inf'_le_iff : s.inf' H f ≤ a ↔ ∃ i ∈ s, f i ≤ a := +@le_sup'_iff (order_dual α) _ _ _ H f _ + +@[simp] lemma inf'_lt_iff : s.inf' H f < a ↔ ∃ i ∈ s, f i < a := +@lt_sup'_iff (order_dual α) _ _ _ H f _ + +@[simp] lemma lt_inf'_iff : a < s.inf' H f ↔ ∀ i ∈ s, a < f i := +@sup'_lt_iff (order_dual α) _ _ _ H f _ + +lemma exists_mem_eq_sup' (f : ι → α) : ∃ i, i ∈ s ∧ s.sup' H f = f i := +begin + refine H.cons_induction (λ c, _) (λ c s hc hs ih, _), + { exact ⟨c, mem_singleton_self c, rfl⟩, }, + { rcases ih with ⟨b, hb, h'⟩, + rw [sup'_cons hs, h'], + cases total_of (≤) (f b) (f c) with h h, + { exact ⟨c, mem_cons.2 (or.inl rfl), sup_eq_left.2 h⟩, }, + { exact ⟨b, mem_cons.2 (or.inr hb), sup_eq_right.2 h⟩, }, }, +end + +lemma exists_mem_eq_inf' (f : ι → α) : ∃ i, i ∈ s ∧ s.inf' H f = f i := +@exists_mem_eq_sup' (order_dual α) _ _ _ H f + +lemma exists_mem_eq_sup [order_bot α] (s : finset ι) (h : s.nonempty) (f : ι → α) : + ∃ i, i ∈ s ∧ s.sup f = f i := +sup'_eq_sup h f ▸ exists_mem_eq_sup' h f + +lemma exists_mem_eq_inf [order_top α] (s : finset ι) (h : s.nonempty) (f : ι → α) : + ∃ i, i ∈ s ∧ s.inf f = f i := +@exists_mem_eq_sup (order_dual α) _ _ _ _ h f + +end linear_order + /-! ### max and min of finite sets -/ section max_min variables [linear_order α] diff --git a/src/data/real/nnreal.lean b/src/data/real/nnreal.lean index 2c66a26ed9684..59946ac4fb2c4 100644 --- a/src/data/real/nnreal.lean +++ b/src/data/real/nnreal.lean @@ -414,7 +414,7 @@ by simp [nnreal.coe_le_coe.symm, real.to_nnreal, hp] @[simp] lemma to_nnreal_lt_to_nnreal_iff' {r p : ℝ} : real.to_nnreal r < real.to_nnreal p ↔ r < p ∧ 0 < p := -by simp [nnreal.coe_lt_coe.symm, real.to_nnreal, lt_irrefl] +nnreal.coe_lt_coe.symm.trans max_lt_max_left_iff lemma to_nnreal_lt_to_nnreal_iff {r p : ℝ} (h : 0 < p) : real.to_nnreal r < real.to_nnreal p ↔ r < p := diff --git a/src/data/set/intervals/basic.lean b/src/data/set/intervals/basic.lean index ce8edd65c4dd7..c4a7a2815a8a5 100644 --- a/src/data/set/intervals/basic.lean +++ b/src/data/set/intervals/basic.lean @@ -25,7 +25,7 @@ for some statements it should be `linear_order` or `densely_ordered`). TODO: This is just the beginning; a lot of rules are missing -/ -universe u +variables {α β : Type*} namespace set @@ -33,7 +33,7 @@ open set open order_dual (to_dual of_dual) section preorder -variables {α : Type u} [preorder α] {a a₁ a₂ b b₁ b₂ c x : α} +variables [preorder α] {a a₁ a₂ b b₁ b₂ c x : α} /-- Left-open right-open interval -/ def Ioo (a b : α) := {x | a < x ∧ x < b} @@ -385,7 +385,7 @@ ext $ λ x, ⟨λ H, ⟨H.2.1, H.1⟩, λ H, ⟨H.2, H.1, H.2.trans h⟩⟩ end preorder section partial_order -variables {α : Type u} [partial_order α] {a b c : α} +variables [partial_order α] {a b c : α} @[simp] lemma Icc_self (a : α) : Icc a a = {a} := set.ext $ by simp [Icc, le_antisymm_iff, and_comm] @@ -529,10 +529,9 @@ end partial_order section order_top -@[simp] lemma Ici_top {α : Type u} [partial_order α] [order_top α] : Ici (⊤ : α) = {⊤} := -is_max_top.Ici_eq +@[simp] lemma Ici_top [partial_order α] [order_top α] : Ici (⊤ : α) = {⊤} := is_max_top.Ici_eq -variables {α : Type u} [preorder α] [order_top α] {a : α} +variables [preorder α] [order_top α] {a : α} @[simp] lemma Ioi_top : Ioi (⊤ : α) = ∅ := is_max_top.Ioi_eq @[simp] lemma Iic_top : Iic (⊤ : α) = univ := is_top_top.Iic_eq @@ -543,10 +542,10 @@ end order_top section order_bot -@[simp] lemma Iic_bot {α : Type u} [partial_order α] [order_bot α] : Iic (⊥ : α) = {⊥} := +@[simp] lemma Iic_bot [partial_order α] [order_bot α] : Iic (⊥ : α) = {⊥} := is_min_bot.Iic_eq -variables {α : Type u} [preorder α] [order_bot α] {a : α} +variables [preorder α] [order_bot α] {a : α} @[simp] lemma Iio_bot : Iio (⊥ : α) = ∅ := is_min_bot.Iio_eq @[simp] lemma Ici_bot : Ici (⊥ : α) = univ := is_bot_bot.Ici_eq @@ -555,8 +554,10 @@ variables {α : Type u} [preorder α] [order_bot α] {a : α} end order_bot +lemma Icc_bot_top [partial_order α] [bounded_order α] : Icc (⊥ : α) ⊤ = univ := by simp + section linear_order -variables {α : Type u} [linear_order α] {a a₁ a₂ b b₁ b₂ c d : α} +variables [linear_order α] {a a₁ a₂ b b₁ b₂ c d : α} lemma not_mem_Ici : c ∉ Ici a ↔ c < a := not_le @@ -1132,14 +1133,11 @@ section lattice section inf -variables {α : Type u} [semilattice_inf α] +variables [semilattice_inf α] @[simp] lemma Iic_inter_Iic {a b : α} : Iic a ∩ Iic b = Iic (a ⊓ b) := by { ext x, simp [Iic] } -@[simp] lemma Iio_inter_Iio [is_total α (≤)] {a b : α} : Iio a ∩ Iio b = Iio (a ⊓ b) := -by { ext x, simp [Iio] } - @[simp] lemma Ioc_inter_Iic (a b c : α) : Ioc a b ∩ Iic c = Ioc a (b ⊓ c) := by rw [← Ioi_inter_Iic, ← Ioi_inter_Iic, inter_assoc, Iic_inter_Iic] @@ -1147,7 +1145,7 @@ end inf section sup -variables {α : Type u} [semilattice_sup α] +variables [semilattice_sup α] @[simp] lemma Ici_inter_Ici {a b : α} : Ici a ∩ Ici b = Ici (a ⊔ b) := by { ext x, simp [Ici] } @@ -1155,18 +1153,11 @@ by { ext x, simp [Ici] } @[simp] lemma Ico_inter_Ici (a b c : α) : Ico a b ∩ Ici c = Ico (a ⊔ c) b := by rw [← Ici_inter_Iio, ← Ici_inter_Iio, ← Ici_inter_Ici, inter_right_comm] -@[simp] lemma Ioi_inter_Ioi [is_total α (≤)] {a b : α} : Ioi a ∩ Ioi b = Ioi (a ⊔ b) := -by { ext x, simp [Ioi] } - -@[simp] lemma Ioc_inter_Ioi [is_total α (≤)] {a b c : α} : Ioc a b ∩ Ioi c = Ioc (a ⊔ c) b := -by rw [← Ioi_inter_Iic, inter_assoc, inter_comm, inter_assoc, Ioi_inter_Ioi, inter_comm, - Ioi_inter_Iic, sup_comm] - end sup section both -variables {α : Type u} [lattice α] [ht : is_total α (≤)] {a b c a₁ a₂ b₁ b₂ : α} +variables [lattice α] {a b c a₁ a₂ b₁ b₂ : α} lemma Icc_inter_Icc : Icc a₁ b₁ ∩ Icc a₂ b₂ = Icc (a₁ ⊔ a₂) (b₁ ⊓ b₂) := by simp only [Ici_inter_Iic.symm, Ici_inter_Ici.symm, Iic_inter_Iic.symm]; ac_refl @@ -1175,7 +1166,14 @@ by simp only [Ici_inter_Iic.symm, Ici_inter_Ici.symm, Iic_inter_Iic.symm]; ac_re Icc a b ∩ Icc b c = {b} := by rw [Icc_inter_Icc, sup_of_le_right hab, inf_of_le_left hbc, Icc_self] -include ht +end both +end lattice + +section linear_order +variables [linear_order α] {a a₁ a₂ b b₁ b₂ c d : α} + +@[simp] lemma Ioi_inter_Ioi : Ioi a ∩ Ioi b = Ioi (a ⊔ b) := ext $ λ _, sup_lt_iff.symm +@[simp] lemma Iio_inter_Iio : Iio a ∩ Iio b = Iio (a ⊓ b) := ext $ λ _, lt_inf_iff.symm lemma Ico_inter_Ico : Ico a₁ b₁ ∩ Ico a₂ b₂ = Ico (a₁ ⊔ a₂) (b₁ ⊓ b₂) := by simp only [Ici_inter_Iio.symm, Ici_inter_Ici.symm, Iio_inter_Iio.symm]; ac_refl @@ -1186,15 +1184,6 @@ by simp only [Ioi_inter_Iic.symm, Ioi_inter_Ioi.symm, Iic_inter_Iic.symm]; ac_re lemma Ioo_inter_Ioo : Ioo a₁ b₁ ∩ Ioo a₂ b₂ = Ioo (a₁ ⊔ a₂) (b₁ ⊓ b₂) := by simp only [Ioi_inter_Iio.symm, Ioi_inter_Ioi.symm, Iio_inter_Iio.symm]; ac_refl -end both - -lemma Icc_bot_top {α} [partial_order α] [bounded_order α] : Icc (⊥ : α) ⊤ = univ := by simp - -end lattice - -section linear_order -variables {α : Type u} [linear_order α] {a a₁ a₂ b b₁ b₂ c d : α} - lemma Ioc_inter_Ioo_of_left_lt (h : b₁ < b₂) : Ioc a₁ b₁ ∩ Ioo a₂ b₂ = Ioc (max a₁ a₂) b₁ := ext $ λ x, by simp [and_assoc, @and.left_comm (x ≤ _), and_iff_left_iff_imp.2 (λ h', lt_of_le_of_lt h' h)] @@ -1215,6 +1204,10 @@ by rw [diff_eq, compl_Iio, Ico_inter_Ici, sup_eq_max] @[simp] lemma Ioc_diff_Ioi : Ioc a b \ Ioi c = Ioc a (min b c) := ext $ by simp [iff_def] {contextual:=tt} +@[simp] lemma Ioc_inter_Ioi : Ioc a b ∩ Ioi c = Ioc (a ⊔ c) b := +by rw [← Ioi_inter_Iic, inter_assoc, inter_comm, inter_assoc, Ioi_inter_Ioi, inter_comm, + Ioi_inter_Iic, sup_comm] + @[simp] lemma Ico_inter_Iio : Ico a b ∩ Iio c = Ico a (min b c) := ext $ by simp [iff_def] {contextual:=tt} @@ -1247,7 +1240,7 @@ end linear_order section prod -variables {α β : Type*} [preorder α] [preorder β] +variables [preorder α] [preorder β] @[simp] lemma Iic_prod_Iic (a : α) (b : β) : Iic a ×ˢ Iic b = Iic (a, b) := rfl @@ -1271,7 +1264,7 @@ end prod section ordered_comm_group -variables {α : Type*} [ordered_comm_group α] {a b c d : α} +variables [ordered_comm_group α] {a b c d : α} /-! `inv_mem_Ixx_iff`, `sub_mem_Ixx_iff` -/ @[to_additive] lemma inv_mem_Icc_iff : a⁻¹ ∈ set.Icc c d ↔ a ∈ set.Icc (d⁻¹) (c⁻¹) := @@ -1287,7 +1280,7 @@ end ordered_comm_group section ordered_add_comm_group -variables {α : Type*} [ordered_add_comm_group α] {a b c d : α} +variables [ordered_add_comm_group α] {a b c d : α} /-! `add_mem_Ixx_iff_left` -/ lemma add_mem_Icc_iff_left : a + b ∈ set.Icc c d ↔ a ∈ set.Icc (c - b) (d - b) := @@ -1339,7 +1332,7 @@ end ordered_add_comm_group section linear_ordered_add_comm_group -variables {α : Type u} [linear_ordered_add_comm_group α] +variables [linear_ordered_add_comm_group α] /-- If we remove a smaller interval from a larger, the result is nonempty -/ lemma nonempty_Ico_sdiff {x dx y dy : α} (h : dy < dx) (hx : 0 < dx) : @@ -1357,7 +1350,6 @@ end set open set namespace order_iso -variables {α β : Type*} section preorder variables [preorder α] [preorder β] @@ -1428,7 +1420,7 @@ end order_iso section dense -variables (α : Type*) [preorder α] [densely_ordered α] {x y : α} +variables (α) [preorder α] [densely_ordered α] {x y : α} instance : no_min_order (set.Ioo x y) := ⟨λ ⟨a, ha₁, ha₂⟩, begin diff --git a/src/order/complete_lattice.lean b/src/order/complete_lattice.lean index 5ab8354a8a17f..c299d80d2ebb6 100644 --- a/src/order/complete_lattice.lean +++ b/src/order/complete_lattice.lean @@ -1177,6 +1177,9 @@ instance Prop.complete_lattice : complete_lattice Prop := .. Prop.bounded_order, .. Prop.distrib_lattice } +noncomputable instance Prop.complete_linear_order : complete_linear_order Prop := +{ ..Prop.complete_lattice, ..Prop.linear_order } + @[simp] lemma Sup_Prop_eq {s : set Prop} : Sup s = ∃ p ∈ s, p := rfl @[simp] lemma Inf_Prop_eq {s : set Prop} : Inf s = ∀ p ∈ s, p := rfl diff --git a/src/order/lattice.lean b/src/order/lattice.lean index 290472351972a..a0aeec33efa9e 100644 --- a/src/order/lattice.lean +++ b/src/order/lattice.lean @@ -186,24 +186,6 @@ sup_le_sup h₁ le_rfl theorem le_of_sup_eq (h : a ⊔ b = b) : a ≤ b := by { rw ← h, simp } -lemma sup_ind [is_total α (≤)] (a b : α) {p : α → Prop} (ha : p a) (hb : p b) : p (a ⊔ b) := -(is_total.total a b).elim (λ h : a ≤ b, by rwa sup_eq_right.2 h) (λ h, by rwa sup_eq_left.2 h) - -@[simp] lemma sup_lt_iff [is_total α (≤)] {a b c : α} : b ⊔ c < a ↔ b < a ∧ c < a := -⟨λ h, ⟨le_sup_left.trans_lt h, le_sup_right.trans_lt h⟩, λ h, sup_ind b c h.1 h.2⟩ - -@[simp] lemma le_sup_iff [is_total α (≤)] {a b c : α} : a ≤ b ⊔ c ↔ a ≤ b ∨ a ≤ c := -⟨λ h, (total_of (≤) c b).imp - (λ bc, by rwa sup_eq_left.2 bc at h) - (λ bc, by rwa sup_eq_right.2 bc at h), - λ h, h.elim le_sup_of_le_left le_sup_of_le_right⟩ - -@[simp] lemma lt_sup_iff [is_total α (≤)] {a b c : α} : a < b ⊔ c ↔ a < b ∨ a < c := -⟨λ h, (total_of (≤) c b).imp - (λ bc, by rwa sup_eq_left.2 bc at h) - (λ bc, by rwa sup_eq_right.2 bc at h), - λ h, h.elim lt_sup_of_lt_left lt_sup_of_lt_right⟩ - @[simp] theorem sup_idem : a ⊔ a = a := by apply le_antisymm; simp @@ -381,15 +363,6 @@ inf_le_inf le_rfl h theorem le_of_inf_eq (h : a ⊓ b = a) : a ≤ b := by { rw ← h, simp } -lemma inf_ind [is_total α (≤)] (a b : α) {p : α → Prop} (ha : p a) (hb : p b) : p (a ⊓ b) := -@sup_ind (order_dual α) _ _ _ _ _ ha hb - -@[simp] lemma lt_inf_iff [is_total α (≤)] {a b c : α} : a < b ⊓ c ↔ a < b ∧ a < c := -@sup_lt_iff (order_dual α) _ _ _ _ _ - -@[simp] lemma inf_le_iff [is_total α (≤)] {a b c : α} : b ⊓ c ≤ a ↔ b ≤ a ∨ c ≤ a := -@le_sup_iff (order_dual α) _ _ _ _ _ - @[simp] theorem inf_idem : a ⊓ a = a := @sup_idem (order_dual α) _ _ @@ -666,8 +639,38 @@ instance linear_order.to_lattice {α : Type u} [o : linear_order α] : le_inf := assume a b c, le_min, ..o } -theorem sup_eq_max [linear_order α] {x y : α} : x ⊔ y = max x y := rfl -theorem inf_eq_min [linear_order α] {x y : α} : x ⊓ y = min x y := rfl +section linear_order +variables [linear_order α] {a b c : α} + +lemma sup_eq_max : a ⊔ b = max a b := rfl +lemma inf_eq_min : a ⊓ b = min a b := rfl + +lemma sup_ind (a b : α) {p : α → Prop} (ha : p a) (hb : p b) : p (a ⊔ b) := +(is_total.total a b).elim (λ h : a ≤ b, by rwa sup_eq_right.2 h) (λ h, by rwa sup_eq_left.2 h) + +@[simp] lemma le_sup_iff : a ≤ b ⊔ c ↔ a ≤ b ∨ a ≤ c := +⟨λ h, (total_of (≤) c b).imp + (λ bc, by rwa sup_eq_left.2 bc at h) + (λ bc, by rwa sup_eq_right.2 bc at h), + λ h, h.elim le_sup_of_le_left le_sup_of_le_right⟩ + +@[simp] lemma lt_sup_iff : a < b ⊔ c ↔ a < b ∨ a < c := +⟨λ h, (total_of (≤) c b).imp + (λ bc, by rwa sup_eq_left.2 bc at h) + (λ bc, by rwa sup_eq_right.2 bc at h), + λ h, h.elim lt_sup_of_lt_left lt_sup_of_lt_right⟩ + +@[simp] lemma sup_lt_iff : b ⊔ c < a ↔ b < a ∧ c < a := +⟨λ h, ⟨le_sup_left.trans_lt h, le_sup_right.trans_lt h⟩, λ h, sup_ind b c h.1 h.2⟩ + +lemma inf_ind (a b : α) {p : α → Prop} (ha : p a) (hb : p b) : p (a ⊓ b) := +@sup_ind (order_dual α) _ _ _ _ ha hb + +@[simp] lemma inf_le_iff : b ⊓ c ≤ a ↔ b ≤ a ∨ c ≤ a := @le_sup_iff (order_dual α) _ _ _ _ +@[simp] lemma inf_lt_iff : b ⊓ c < a ↔ b < a ∨ c < a := @lt_sup_iff (order_dual α) _ _ _ _ +@[simp] lemma lt_inf_iff : a < b ⊓ c ↔ a < b ∧ a < c := @sup_lt_iff (order_dual α) _ _ _ _ + +end linear_order lemma sup_eq_max_default [semilattice_sup α] [decidable_rel ((≤) : α → α → Prop)] [is_total α (≤)] : (⊔) = (max_default : α → α → α) := @@ -773,22 +776,20 @@ lemma le_map_sup [semilattice_sup α] [semilattice_sup β] f x ⊔ f y ≤ f (x ⊔ y) := sup_le (h le_sup_left) (h le_sup_right) -lemma map_sup [semilattice_sup α] [is_total α (≤)] [semilattice_sup β] {f : α → β} - (hf : monotone f) (x y : α) : - f (x ⊔ y) = f x ⊔ f y := -(is_total.total x y).elim - (λ h : x ≤ y, by simp only [h, hf h, sup_of_le_right]) - (λ h, by simp only [h, hf h, sup_of_le_left]) - lemma map_inf_le [semilattice_inf α] [semilattice_inf β] {f : α → β} (h : monotone f) (x y : α) : f (x ⊓ y) ≤ f x ⊓ f y := le_inf (h inf_le_left) (h inf_le_right) -lemma map_inf [semilattice_inf α] [is_total α (≤)] [semilattice_inf β] {f : α → β} - (hf : monotone f) (x y : α) : - f (x ⊓ y) = f x ⊓ f y := -@monotone.map_sup (order_dual α) _ _ _ _ _ hf.dual x y +variables [linear_order α] + +lemma map_sup [semilattice_sup β] {f : α → β} (hf : monotone f) (x y : α) : f (x ⊔ y) = f x ⊔ f y := +(is_total.total x y).elim + (λ h : x ≤ y, by simp only [h, hf h, sup_of_le_right]) + (λ h, by simp only [h, hf h, sup_of_le_left]) + +lemma map_inf [semilattice_inf β] {f : α → β} (hf : monotone f) (x y : α) : f (x ⊓ y) = f x ⊓ f y := +hf.dual.map_sup _ _ end monotone @@ -819,19 +820,17 @@ lemma map_sup_le [semilattice_sup α] [semilattice_inf β] f (x ⊔ y) ≤ f x ⊓ f y := h.dual_right.le_map_sup x y -lemma map_sup [semilattice_sup α] [is_total α (≤)] [semilattice_inf β] {f : α → β} - (hf : antitone f) (x y : α) : - f (x ⊔ y) = f x ⊓ f y := -hf.dual_right.map_sup x y - lemma le_map_inf [semilattice_inf α] [semilattice_sup β] {f : α → β} (h : antitone f) (x y : α) : f x ⊔ f y ≤ f (x ⊓ y) := h.dual_right.map_inf_le x y -lemma map_inf [semilattice_inf α] [is_total α (≤)] [semilattice_sup β] {f : α → β} - (hf : antitone f) (x y : α) : - f (x ⊓ y) = f x ⊔ f y := +variables [linear_order α] + +lemma map_sup [semilattice_inf β] {f : α → β} (hf : antitone f) (x y : α) : f (x ⊔ y) = f x ⊓ f y := +hf.dual_right.map_sup x y + +lemma map_inf [semilattice_sup β] {f : α → β} (hf : antitone f) (x y : α) : f (x ⊓ y) = f x ⊔ f y := hf.dual_right.map_inf x y end antitone diff --git a/src/order/min_max.lean b/src/order/min_max.lean index 73f90ab750dd5..cbe72653f0ef4 100644 --- a/src/order/min_max.lean +++ b/src/order/min_max.lean @@ -25,7 +25,13 @@ variables [linear_order α] [linear_order β] {f : α → β} {s : set α} {a b -- translate from lattices to linear orders (sup → max, inf → min) @[simp] lemma le_min_iff : c ≤ min a b ↔ c ≤ a ∧ c ≤ b := le_inf_iff +@[simp] lemma le_max_iff : a ≤ max b c ↔ a ≤ b ∨ a ≤ c := le_sup_iff +@[simp] lemma min_le_iff : min a b ≤ c ↔ a ≤ c ∨ b ≤ c := inf_le_iff @[simp] lemma max_le_iff : max a b ≤ c ↔ a ≤ c ∧ b ≤ c := sup_le_iff +@[simp] lemma lt_min_iff : a < min b c ↔ a < b ∧ a < c := lt_inf_iff +@[simp] lemma lt_max_iff : a < max b c ↔ a < b ∨ a < c := lt_sup_iff +@[simp] lemma min_lt_iff : min a b < c ↔ a < c ∨ b < c := inf_lt_iff +@[simp] lemma max_lt_iff : max a b < c ↔ a < c ∧ b < c := sup_lt_iff lemma max_le_max : a ≤ c → b ≤ d → max a b ≤ max c d := sup_le_sup lemma min_le_min : a ≤ c → b ≤ d → min a b ≤ min c d := inf_le_inf lemma le_max_of_le_left : a ≤ b → a ≤ max b c := le_sup_of_le_left @@ -78,30 +84,24 @@ end lemma max_eq_iff : max a b = c ↔ a = c ∧ b ≤ a ∨ b = c ∧ a ≤ b := @min_eq_iff (order_dual α) _ a b c +lemma min_lt_min_left_iff : min a c < min b c ↔ a < b ∧ a < c := +by { simp_rw [lt_min_iff, min_lt_iff, or_iff_left (lt_irrefl _)], + exact and_congr_left (λ h, or_iff_left_of_imp h.trans) } + +lemma min_lt_min_right_iff : min a b < min a c ↔ b < c ∧ b < a := +by simp_rw [min_comm a, min_lt_min_left_iff] + +lemma max_lt_max_left_iff : max a c < max b c ↔ a < b ∧ c < b := +@min_lt_min_left_iff (order_dual α) _ _ _ _ +lemma max_lt_max_right_iff : max a b < max a c ↔ b < c ∧ a < c := +@min_lt_min_right_iff (order_dual α) _ _ _ _ + /-- An instance asserting that `max a a = a` -/ instance max_idem : is_idempotent α max := by apply_instance -- short-circuit type class inference /-- An instance asserting that `min a a = a` -/ instance min_idem : is_idempotent α min := by apply_instance -- short-circuit type class inference -@[simp] lemma max_lt_iff : max a b < c ↔ (a < c ∧ b < c) := -sup_lt_iff - -@[simp] lemma lt_min_iff : a < min b c ↔ (a < b ∧ a < c) := -lt_inf_iff - -@[simp] lemma lt_max_iff : a < max b c ↔ a < b ∨ a < c := -lt_sup_iff - -@[simp] lemma min_lt_iff : min a b < c ↔ a < c ∨ b < c := -@lt_max_iff (order_dual α) _ _ _ _ - -@[simp] lemma min_le_iff : min a b ≤ c ↔ a ≤ c ∨ b ≤ c := -inf_le_iff - -@[simp] lemma le_max_iff : a ≤ max b c ↔ a ≤ b ∨ a ≤ c := -@min_le_iff (order_dual α) _ _ _ _ - lemma min_lt_max : min a b < max a b ↔ a ≠ b := inf_lt_sup lemma max_lt_max (h₁ : a < c) (h₂ : b < d) : max a b < max c d := diff --git a/src/order/omega_complete_partial_order.lean b/src/order/omega_complete_partial_order.lean index 741a53fb83d85..8a32d6fda2117 100644 --- a/src/order/omega_complete_partial_order.lean +++ b/src/order/omega_complete_partial_order.lean @@ -388,6 +388,8 @@ instance : omega_complete_partial_order (α × β) := end prod +open omega_complete_partial_order + namespace complete_lattice variables (α : Type u) @@ -402,26 +404,6 @@ instance [complete_lattice α] : omega_complete_partial_order α := le_ωSup := assume ⟨c, _⟩ i, by simp only [order_hom.coe_fun_mk]; apply le_supr_of_le i; refl } variables {α} {β : Type v} [omega_complete_partial_order α] [complete_lattice β] -open omega_complete_partial_order - -lemma inf_continuous [is_total β (≤)] (f g : α →o β) (hf : continuous f) (hg : continuous g) : - continuous (f ⊓ g) := -begin - intro c, - apply eq_of_forall_ge_iff, intro z, - simp only [inf_le_iff, hf c, hg c, ωSup_le_iff, ←forall_or_distrib_left, ←forall_or_distrib_right, - function.comp_app, chain.map_coe, order_hom.has_inf_inf_coe], - split, - { introv h, apply h }, - { intros h i j, - apply or.imp _ _ (h (max i j)); apply le_trans; mono*; try { exact le_rfl }, - { apply le_max_left }, - { apply le_max_right }, }, -end - -lemma inf_continuous' [is_total β (≤)] {f g : α → β} (hf : continuous' f) (hg : continuous' g) : - continuous' (f ⊓ g) := -⟨_, inf_continuous _ _ hf.snd hg.snd⟩ lemma Sup_continuous (s : set $ α →o β) (hs : ∀ f ∈ s, continuous f) : continuous (Sup s) := @@ -470,6 +452,24 @@ end end complete_lattice +namespace complete_lattice +variables {α β : Type*} [omega_complete_partial_order α] [complete_linear_order β] + +lemma inf_continuous (f g : α →o β) (hf : continuous f) (hg : continuous g) : continuous (f ⊓ g) := +begin + refine λ c, eq_of_forall_ge_iff (λ z, _), + simp only [inf_le_iff, hf c, hg c, ωSup_le_iff, ←forall_or_distrib_left, ←forall_or_distrib_right, + function.comp_app, chain.map_coe, order_hom.has_inf_inf_coe], + exact ⟨λ h _, h _ _, λ h i j, (h (max i j)).imp (le_trans $ f.mono $ c.mono $ le_max_left _ _) + (le_trans $ g.mono $ c.mono $ le_max_right _ _)⟩, +end + +lemma inf_continuous' {f g : α → β} (hf : continuous' f) (hg : continuous' g) : + continuous' (f ⊓ g) := +⟨_, inf_continuous _ _ hf.snd hg.snd⟩ + +end complete_lattice + namespace omega_complete_partial_order variables {α : Type u} {α' : Type*} {β : Type v} {β' : Type*} {γ : Type*} {φ : Type*} diff --git a/src/topology/algebra/order/basic.lean b/src/topology/algebra/order/basic.lean index d90d7a85bcb9c..f1cb85b76af35 100644 --- a/src/topology/algebra/order/basic.lean +++ b/src/topology/algebra/order/basic.lean @@ -927,8 +927,8 @@ lemma nhds_bot_order [topological_space α] [partial_order α] [order_bot α] [o 𝓝 (⊥:α) = (⨅l (h₂ : ⊥ < l), 𝓟 (Iio l)) := by simp [nhds_eq_order (⊥:α)] -lemma nhds_top_basis [topological_space α] [semilattice_sup α] [order_top α] [is_total α has_le.le] - [order_topology α] [nontrivial α] : +lemma nhds_top_basis [topological_space α] [linear_order α] [order_top α] [order_topology α] + [nontrivial α] : (𝓝 ⊤).has_basis (λ a : α, a < ⊤) (λ a : α, Ioi a) := ⟨ begin simp only [nhds_top_order], @@ -941,22 +941,22 @@ lemma nhds_top_basis [topological_space α] [semilattice_sup α] [order_top α] exact ⟨a, lt_top_iff_ne_top.mpr ha⟩ } end ⟩ -lemma nhds_bot_basis [topological_space α] [semilattice_inf α] [order_bot α] [is_total α has_le.le] - [order_topology α] [nontrivial α] : +lemma nhds_bot_basis [topological_space α] [linear_order α] [order_bot α] [order_topology α] + [nontrivial α] : (𝓝 ⊥).has_basis (λ a : α, ⊥ < a) (λ a : α, Iio a) := -@nhds_top_basis (order_dual α) _ _ _ _ _ _ +@nhds_top_basis (order_dual α) _ _ _ _ _ -lemma nhds_top_basis_Ici [topological_space α] [semilattice_sup α] [order_top α] - [is_total α has_le.le] [order_topology α] [nontrivial α] [densely_ordered α] : +lemma nhds_top_basis_Ici [topological_space α] [linear_order α] [order_top α] [order_topology α] + [nontrivial α] [densely_ordered α] : (𝓝 ⊤).has_basis (λ a : α, a < ⊤) Ici := nhds_top_basis.to_has_basis (λ a ha, let ⟨b, hab, hb⟩ := exists_between ha in ⟨b, hb, Ici_subset_Ioi.mpr hab⟩) (λ a ha, ⟨a, ha, Ioi_subset_Ici_self⟩) -lemma nhds_bot_basis_Iic [topological_space α] [semilattice_inf α] [order_bot α] - [is_total α has_le.le] [order_topology α] [nontrivial α] [densely_ordered α] : +lemma nhds_bot_basis_Iic [topological_space α] [linear_order α] [order_bot α] [order_topology α] + [nontrivial α] [densely_ordered α] : (𝓝 ⊥).has_basis (λ a : α, ⊥ < a) Iic := -@nhds_top_basis_Ici (order_dual α) _ _ _ _ _ _ _ +@nhds_top_basis_Ici (order_dual α) _ _ _ _ _ _ lemma tendsto_nhds_top_mono [topological_space β] [partial_order β] [order_top β] [order_topology β] {l : filter α} {f g : α → β} (hf : tendsto f l (𝓝 ⊤)) (hg : f ≤ᶠ[l] g) : From 82b9c42f7adbe4e0a494fa0cbdd2651d0f49c597 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 3 May 2022 12:18:54 +0000 Subject: [PATCH 037/153] feat(set_theory/game/nim): Mark many lemmas as `simp` (#13844) --- src/set_theory/game/nim.lean | 47 ++++++++++++++++++------------------ 1 file changed, 24 insertions(+), 23 deletions(-) diff --git a/src/set_theory/game/nim.lean b/src/set_theory/game/nim.lean index e2d6b5099b84c..cf8bb340f12f5 100644 --- a/src/set_theory/game/nim.lean +++ b/src/set_theory/game/nim.lean @@ -90,7 +90,7 @@ by { rw nim_def, exact λ i, ⟨_, ⟨ordinal.typein_lt_self i, rfl⟩⟩ } lemma exists_move_left_eq {O : ordinal} : ∀ O' < O, ∃ i, (nim O).move_left i = nim O' := by { rw nim_def, exact λ _ h, ⟨(ordinal.principal_seg_out h).top, by simp⟩ } -lemma zero_first_loses : (nim (0 : ordinal)).first_loses := +@[simp] lemma zero_first_loses : (nim (0 : ordinal)).first_loses := begin rw [impartial.first_loses_symm, nim_def, le_def_lt], exact ⟨@is_empty_elim (0 : ordinal).out.α _ _, @is_empty_elim pempty _ _⟩ @@ -103,7 +103,7 @@ begin exact or.inr ⟨(ordinal.principal_seg_out hO).top, by simpa using zero_first_loses.1⟩ end -lemma sum_first_loses_iff_eq (O₁ O₂ : ordinal) : (nim O₁ + nim O₂).first_loses ↔ O₁ = O₂ := +@[simp] lemma sum_first_loses_iff_eq (O₁ O₂ : ordinal) : (nim O₁ + nim O₂).first_loses ↔ O₁ = O₂ := begin split, { contrapose, @@ -121,10 +121,10 @@ begin exact impartial.add_self (nim O₁) } end -lemma sum_first_wins_iff_neq (O₁ O₂ : ordinal) : (nim O₁ + nim O₂).first_wins ↔ O₁ ≠ O₂ := +@[simp] lemma sum_first_wins_iff_neq (O₁ O₂ : ordinal) : (nim O₁ + nim O₂).first_wins ↔ O₁ ≠ O₂ := by rw [iff_not_comm, impartial.not_first_wins, sum_first_loses_iff_eq] -lemma equiv_iff_eq (O₁ O₂ : ordinal) : nim O₁ ≈ nim O₂ ↔ O₁ = O₂ := +@[simp] lemma equiv_iff_eq (O₁ O₂ : ordinal) : nim O₁ ≈ nim O₂ ↔ O₁ = O₂ := ⟨λ h, (sum_first_loses_iff_eq _ _).1 $ by rw [first_loses_of_equiv_iff (add_congr h (equiv_refl _)), sum_first_loses_iff_eq], by { rintro rfl, refl }⟩ @@ -133,17 +133,17 @@ end nim /-- The Grundy value of an impartial game, the ordinal which corresponds to the game of nim that the game is equivalent to -/ -noncomputable def grundy_value : Π (G : pgame.{u}) [G.impartial], ordinal.{u} -| G := λ hG, by exactI ordinal.mex.{u u} (λ i, grundy_value (G.move_left i)) +noncomputable def grundy_value : Π (G : pgame.{u}), ordinal.{u} +| G := ordinal.mex.{u u} (λ i, grundy_value (G.move_left i)) using_well_founded { dec_tac := pgame_wf_tac } -lemma grundy_value_def (G : pgame) [G.impartial] : +lemma grundy_value_def (G : pgame) : grundy_value G = ordinal.mex.{u u} (λ i, grundy_value (G.move_left i)) := by rw grundy_value /-- The Sprague-Grundy theorem which states that every impartial game is equivalent to a game of nim, namely the game of nim corresponding to the games Grundy value -/ -theorem equiv_nim_grundy_value : ∀ (G : pgame.{u}) [G.impartial], by exactI G ≈ nim (grundy_value G) +theorem equiv_nim_grundy_value : ∀ (G : pgame.{u}) [G.impartial], G ≈ nim (grundy_value G) | G := begin introI hG, @@ -182,25 +182,26 @@ begin end using_well_founded { dec_tac := pgame_wf_tac } -lemma equiv_nim_iff_grundy_value_eq (G : pgame) [G.impartial] (O : ordinal) : - G ≈ nim O ↔ grundy_value G = O := -⟨by { intro h, rw ←nim.equiv_iff_eq, exact equiv_trans (equiv_symm (equiv_nim_grundy_value G)) h }, - by { rintro rfl, exact equiv_nim_grundy_value G }⟩ +@[simp] lemma grundy_value_eq_iff_equiv_nim (G : pgame) [G.impartial] (O : ordinal) : + grundy_value G = O ↔ G ≈ nim O := +⟨by { rintro rfl, exact equiv_nim_grundy_value G }, + by { intro h, rw ←nim.equiv_iff_eq, exact equiv_trans (equiv_symm (equiv_nim_grundy_value G)) h }⟩ lemma nim.grundy_value (O : ordinal.{u}) : grundy_value (nim O) = O := -by rw ←equiv_nim_iff_grundy_value_eq +by simp -lemma equiv_iff_grundy_value_eq (G H : pgame) [G.impartial] [H.impartial] : - G ≈ H ↔ grundy_value G = grundy_value H := -(equiv_congr_left.1 (equiv_nim_grundy_value H) _).trans $ equiv_nim_iff_grundy_value_eq _ _ +@[simp] lemma grundy_value_eq_iff_equiv (G H : pgame) [G.impartial] [H.impartial] : + grundy_value G = grundy_value H ↔ G ≈ H := +(grundy_value_eq_iff_equiv_nim _ _).trans (equiv_congr_left.1 (equiv_nim_grundy_value H) _).symm -lemma grundy_value_zero : grundy_value 0 = 0 := -by rw [(equiv_iff_grundy_value_eq 0 (nim 0)).1 (equiv_symm nim.zero_first_loses), nim.grundy_value] +@[simp] lemma grundy_value_zero : grundy_value 0 = 0 := +by rw [(grundy_value_eq_iff_equiv 0 (nim 0)).2 (equiv_symm nim.zero_first_loses), nim.grundy_value] -lemma equiv_zero_iff_grundy_value (G : pgame) [G.impartial] : G ≈ 0 ↔ grundy_value G = 0 := -by rw [equiv_iff_grundy_value_eq, grundy_value_zero] +@[simp] lemma grundy_value_iff_equiv_zero (G : pgame) [G.impartial] : grundy_value G = 0 ↔ G ≈ 0 := +by rw [←grundy_value_eq_iff_equiv, grundy_value_zero] -lemma grundy_value_nim_add_nim (n m : ℕ) : grundy_value (nim.{u} n + nim.{u} m) = nat.lxor n m := +@[simp] lemma grundy_value_nim_add_nim (n m : ℕ) : + grundy_value (nim.{u} n + nim.{u} m) = nat.lxor n m := begin induction n using nat.strong_induction_on with n hn generalizing m, induction m using nat.strong_induction_on with m hm, @@ -271,12 +272,12 @@ begin end lemma nim_add_nim_equiv {n m : ℕ} : nim n + nim m ≈ nim (nat.lxor n m) := -by rw [equiv_nim_iff_grundy_value_eq, grundy_value_nim_add_nim] +by rw [←grundy_value_eq_iff_equiv_nim, grundy_value_nim_add_nim] lemma grundy_value_add (G H : pgame) [G.impartial] [H.impartial] {n m : ℕ} (hG : grundy_value G = n) (hH : grundy_value H = m) : grundy_value (G + H) = nat.lxor n m := begin - rw [←nim.grundy_value (nat.lxor n m), ←equiv_iff_grundy_value_eq], + rw [←nim.grundy_value (nat.lxor n m), grundy_value_eq_iff_equiv], refine equiv_trans _ nim_add_nim_equiv, convert add_congr (equiv_nim_grundy_value G) (equiv_nim_grundy_value H); simp only [hG, hH] From 9f818ce0fe16e689dc799808e8dcb81d9d028351 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 3 May 2022 12:18:55 +0000 Subject: [PATCH 038/153] =?UTF-8?q?feat(set=5Ftheory/ordinal=5Fbasic):=20`?= =?UTF-8?q?o.out.=CE=B1`=20is=20equivalent=20to=20the=20ordinals=20below?= =?UTF-8?q?=20`o`=20(#13876)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- src/set_theory/ordinal/basic.lean | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/src/set_theory/ordinal/basic.lean b/src/set_theory/ordinal/basic.lean index 75c0a7769f083..e3b09d9dbd9dc 100644 --- a/src/set_theory/ordinal/basic.lean +++ b/src/set_theory/ordinal/basic.lean @@ -681,6 +681,13 @@ enum_type (principal_seg.of_element r a) let ⟨a, e⟩ := typein_surj r h in by clear _let_match; subst e; rw enum_typein +/-- The equivalence between ordinals less than `o` and `o.out.α`. -/ +@[simps] noncomputable def out_equiv_lt (o : ordinal) : {o' : ordinal // o' < o} ≃ o.out.α := +{ to_fun := λ ⟨o', h⟩, enum (<) o' (by rwa type_lt), + inv_fun := λ x, ⟨typein (<) x, typein_lt_self x⟩, + left_inv := λ ⟨o', h⟩, subtype.ext_val (typein_enum _ _), + right_inv := λ h, enum_typein _ _ } + /-- A well order `r` is order isomorphic to the set of ordinals strictly smaller than the ordinal version of `r`. -/ def typein_iso (r : α → α → Prop) [is_well_order α r] : r ≃r subrel (<) (< type r) := From 78c86e1de7ecba60f9ee8aedf423832e7fff2702 Mon Sep 17 00:00:00 2001 From: Stuart Presnell Date: Tue, 3 May 2022 12:18:56 +0000 Subject: [PATCH 039/153] chore(data/nat/totient): golf three lemmas (#13886) Golf the proofs of `totient_le`, `totient_lt`, and `totient_pos` --- src/data/nat/totient.lean | 18 ++++-------------- 1 file changed, 4 insertions(+), 14 deletions(-) diff --git a/src/data/nat/totient.lean b/src/data/nat/totient.lean index df2269212bea7..3ba8b8d2732e3 100644 --- a/src/data/nat/totient.lean +++ b/src/data/nat/totient.lean @@ -39,19 +39,10 @@ by simp [totient] lemma totient_eq_card_coprime (n : ℕ) : φ n = ((range n).filter n.coprime).card := rfl lemma totient_le (n : ℕ) : φ n ≤ n := -calc totient n ≤ (range n).card : card_filter_le _ _ - ... = n : card_range _ +((range n).card_filter_le _).trans_eq (card_range n) lemma totient_lt (n : ℕ) (hn : 1 < n) : φ n < n := -calc totient n ≤ ((range n).filter (≠ 0)).card : - begin - apply card_le_of_subset (monotone_filter_right _ _), - intros n1 hn1 hn1', - simpa only [hn1', coprime_zero_right, hn.ne'] using hn1, - end -... = n - 1 - : by simp only [filter_ne' (range n) 0, card_erase_of_mem, card_range, pos_of_gt hn, mem_range] -... < n : nat.sub_lt (pos_of_gt hn) zero_lt_one +(card_lt_card (filter_ssubset.2 ⟨0, by simp [hn.ne', pos_of_gt hn]⟩)).trans_eq (card_range n) lemma totient_pos : ∀ {n : ℕ}, 0 < n → 0 < φ n | 0 := dec_trivial @@ -135,9 +126,8 @@ if hmn0 : m * n = 0 haveI : fact (0 < (m * n)) := ⟨nat.pos_of_ne_zero hmn0⟩, haveI : fact (0 < m) := ⟨nat.pos_of_ne_zero $ left_ne_zero_of_mul hmn0⟩, haveI : fact (0 < n) := ⟨nat.pos_of_ne_zero $ right_ne_zero_of_mul hmn0⟩, - rw [← zmod.card_units_eq_totient, ← zmod.card_units_eq_totient, - ← zmod.card_units_eq_totient, fintype.card_congr - (units.map_equiv (zmod.chinese_remainder h).to_mul_equiv).to_equiv, + simp only [← zmod.card_units_eq_totient], + rw [fintype.card_congr (units.map_equiv (zmod.chinese_remainder h).to_mul_equiv).to_equiv, fintype.card_congr (@mul_equiv.prod_units (zmod m) (zmod n) _ _).to_equiv, fintype.card_prod] end From 3b971a7bed3dae6f1bc44717e82d5390d48ba1e7 Mon Sep 17 00:00:00 2001 From: tb65536 Date: Tue, 3 May 2022 12:18:57 +0000 Subject: [PATCH 040/153] feat(data/zmod/basic): Add `zmod.cast_sub_one` (#13889) This PR adds `zmod.cast_sub_one`, an analog of `fin.coe_sub_one`. Unfortunately, the proof is a bit long. But maybe it can be golfed? Co-authored-by: tb65536 --- src/data/zmod/basic.lean | 13 +++++++++++++ 1 file changed, 13 insertions(+) diff --git a/src/data/zmod/basic.lean b/src/data/zmod/basic.lean index accb4e6b89a01..f2b80b731a19e 100644 --- a/src/data/zmod/basic.lean +++ b/src/data/zmod/basic.lean @@ -473,6 +473,19 @@ begin { rw [←nat_cast_val, val_neg_one, nat.cast_succ, add_sub_cancel] }, end +lemma cast_sub_one {R : Type*} [ring R] {n : ℕ} (k : zmod n) : + ((k - 1 : zmod n) : R) = (if k = 0 then n else k) - 1 := +begin + split_ifs with hk, + { rw [hk, zero_sub, zmod.cast_neg_one] }, + { cases n, + { rw [int.cast_sub, int.cast_one] }, + { rw [←zmod.nat_cast_val, zmod.val, fin.coe_sub_one, if_neg], + { rw [nat.cast_sub, nat.cast_one, coe_coe], + rwa [fin.ext_iff, fin.coe_zero, ←ne, ←nat.one_le_iff_ne_zero] at hk }, + { exact hk } } }, +end + lemma nat_coe_zmod_eq_iff (p : ℕ) (n : ℕ) (z : zmod p) [fact (0 < p)] : ↑n = z ↔ ∃ k, n = z.val + p * k := begin From 234b3df25a0f94ff15411b64372cc3fdb0615e03 Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Tue, 3 May 2022 14:29:21 +0000 Subject: [PATCH 041/153] feat(analysis/normed_space): lemmas about continuous bilinear maps (#13522) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * Define `continuous_linear_map.map_sub₂` and friends, similar to the lemmas for `linear_map`. * Rename `continuous_linear_map.map_add₂` to `continuous_linear_map.map_add_add` * Two comments refer to `continuous.comp₂`, which will be added in #13423 (but there is otherwise no dependency on this PR). * Define `precompR` and `precompL`, which will be used to compute the derivative of a convolution. * From the sphere eversion project * Required for convolutions --- src/analysis/analytic/linear.lean | 2 +- .../normed_space/bounded_linear_maps.lean | 66 +++++++++++++++++-- src/analysis/normed_space/operator_norm.lean | 28 ++++++-- .../function/strongly_measurable.lean | 5 ++ 4 files changed, 89 insertions(+), 12 deletions(-) diff --git a/src/analysis/analytic/linear.lean b/src/analysis/analytic/linear.lean index cd667f725283e..7ce5a69a2d8f0 100644 --- a/src/analysis/analytic/linear.lean +++ b/src/analysis/analytic/linear.lean @@ -83,7 +83,7 @@ protected theorem has_fpower_series_on_ball_bilinear (f : E →L[𝕜] F →L[ has_sum := λ y _, (has_sum_nat_add_iff' 3).1 $ begin simp only [finset.sum_range_succ, finset.sum_range_one, prod.fst_add, prod.snd_add, - f.map_add₂], + f.map_add_add], dsimp, simp only [add_comm, sub_self, has_sum_zero] end } diff --git a/src/analysis/normed_space/bounded_linear_maps.lean b/src/analysis/normed_space/bounded_linear_maps.lean index 9a980a1eab7fd..663d70c2e8b10 100644 --- a/src/analysis/normed_space/bounded_linear_maps.lean +++ b/src/analysis/normed_space/bounded_linear_maps.lean @@ -51,14 +51,13 @@ artifact, really. noncomputable theory open_locale classical big_operators topological_space -open filter (tendsto) metric +open filter (tendsto) metric continuous_linear_map variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] {E : Type*} [normed_group E] [normed_space 𝕜 E] {F : Type*} [normed_group F] [normed_space 𝕜 F] {G : Type*} [normed_group G] [normed_space 𝕜 G] - /-- A function `f` satisfies `is_bounded_linear_map 𝕜 f` if it is linear and satisfies the inequality `∥f x∥ ≤ M * ∥x∥` for some positive constant `M`. -/ structure is_bounded_linear_map (𝕜 : Type*) [normed_field 𝕜] @@ -228,6 +227,57 @@ end section bilinear_map +namespace continuous_linear_map + +/-! We prove some computation rules for continuous (semi-)bilinear maps in their first argument. + If `f` is a continuuous bilinear map, to use the corresponding rules for the second argument, use + `(f _).map_add` and similar. + + We have to assume that `F` and `G` are normed spaces in this section, to use + `continuous_linear_map.to_normed_group`, but we don't need to assume this for the first argument + of `f`. +-/ + +variables {R : Type*} +variables {𝕜₂ 𝕜' : Type*} [nondiscrete_normed_field 𝕜'] [nondiscrete_normed_field 𝕜₂] +variables {M : Type*} [topological_space M] +variables {σ₁₂ : 𝕜 →+* 𝕜₂} [ring_hom_isometric σ₁₂] +variables {G' : Type*} [normed_group G'] [normed_space 𝕜₂ G'] [normed_space 𝕜' G'] +variables [smul_comm_class 𝕜₂ 𝕜' G'] + +section semiring +variables [semiring R] [add_comm_monoid M] [module R M] {ρ₁₂ : R →+* 𝕜'} + +lemma map_add₂ (f : M →SL[ρ₁₂] F →SL[σ₁₂] G') (x x' : M) (y : F) : + f (x + x') y = f x y + f x' y := +by rw [f.map_add, add_apply] + +lemma map_zero₂ (f : M →SL[ρ₁₂] F →SL[σ₁₂] G') (y : F) : f 0 y = 0 := +by rw [f.map_zero, zero_apply] + +lemma map_smulₛₗ₂ (f : M →SL[ρ₁₂] F →SL[σ₁₂] G') (c : R) (x : M) (y : F) : + f (c • x) y = ρ₁₂ c • f x y := +by rw [f.map_smulₛₗ, smul_apply] +end semiring + +section ring + +variables [ring R] [add_comm_group M] [module R M] {ρ₁₂ : R →+* 𝕜'} + +lemma map_sub₂ (f : M →SL[ρ₁₂] F →SL[σ₁₂] G') (x x' : M) (y : F) : + f (x - x') y = f x y - f x' y := +by rw [f.map_sub, sub_apply] + +lemma map_neg₂ (f : M →SL[ρ₁₂] F →SL[σ₁₂] G') (x : M) (y : F) : f (- x) y = - f x y := +by rw [f.map_neg, neg_apply] + +end ring + +lemma map_smul₂ (f : E →L[𝕜] F →L[𝕜] G) (c : 𝕜) (x : E) (y : F) : f (c • x) y = c • f x y := +by rw [f.map_smul, smul_apply] + +end continuous_linear_map + variable (𝕜) /-- A map `f : E × F → G` satisfies `is_bounded_bilinear_map 𝕜 f` if it is bilinear and @@ -244,10 +294,10 @@ variable {f : E × F → G} lemma continuous_linear_map.is_bounded_bilinear_map (f : E →L[𝕜] F →L[𝕜] G) : is_bounded_bilinear_map 𝕜 (λ x : E × F, f x.1 x.2) := -{ add_left := λ x₁ x₂ y, by rw [f.map_add, continuous_linear_map.add_apply], - smul_left := λ c x y, by rw [f.map_smul _, continuous_linear_map.smul_apply], +{ add_left := f.map_add₂, + smul_left := f.map_smul₂, add_right := λ x, (f x).map_add, - smul_right := λ c x y, (f x).map_smul c y, + smul_right := λ c x, (f x).map_smul c, bound := ⟨max ∥f∥ 1, zero_lt_one.trans_le (le_max_right _ _), λ x y, (f.le_op_norm₂ x y).trans $ by apply_rules [mul_le_mul_of_nonneg_right, norm_nonneg, le_max_left]⟩ } @@ -278,6 +328,7 @@ calc f (x, y - z) = f (x, y + (-1 : 𝕜) • z) : by simp [sub_eq_add_neg] ... = f (x, y) + (-1 : 𝕜) • f (x, z) : by simp only [h.add_right, h.smul_right] ... = f (x, y) - f (x, z) : by simp [sub_eq_add_neg] +/-- Useful to use together with `continuous.comp₂`. -/ lemma is_bounded_bilinear_map.continuous (h : is_bounded_bilinear_map 𝕜 f) : continuous f := begin @@ -317,6 +368,11 @@ lemma is_bounded_bilinear_map.continuous_right (h : is_bounded_bilinear_map 𝕜 continuous (λe₂, f (e₁, e₂)) := h.continuous.comp (continuous_const.prod_mk continuous_id) +/-- Useful to use together with `continuous.comp₂`. -/ +lemma continuous_linear_map.continuous₂ (f : E →L[𝕜] F →L[𝕜] G) : + continuous (function.uncurry (λ x y, f x y)) := +f.is_bounded_bilinear_map.continuous + lemma is_bounded_bilinear_map.is_bounded_linear_map_left (h : is_bounded_bilinear_map 𝕜 f) (y : F) : is_bounded_linear_map 𝕜 (λ x, f (x, y)) := { map_add := λ x x', h.add_left _ _ _, diff --git a/src/analysis/normed_space/operator_norm.lean b/src/analysis/normed_space/operator_norm.lean index 8f0a962823a4d..d76cfa70fb880 100644 --- a/src/analysis/normed_space/operator_norm.lean +++ b/src/analysis/normed_space/operator_norm.lean @@ -26,13 +26,12 @@ noncomputable theory open_locale classical nnreal topological_space -- the `ₗ` subscript variables are for special cases about linear (as opposed to semilinear) maps -variables {𝕜 : Type*} {𝕜₂ : Type*} {𝕜₃ : Type*} {E : Type*} {F : Type*} {Fₗ : Type*} {G : Type*} - {Gₗ : Type*} +variables {𝕜 𝕜₂ 𝕜₃ E Eₗ F Fₗ G Gₗ : Type*} section semi_normed -variables [semi_normed_group E] [semi_normed_group F] [semi_normed_group Fₗ] [semi_normed_group G] - [semi_normed_group Gₗ] +variables [semi_normed_group E] [semi_normed_group Eₗ] [semi_normed_group F] [semi_normed_group Fₗ] +variables [semi_normed_group G] [semi_normed_group Gₗ] open metric continuous_linear_map @@ -126,7 +125,7 @@ rfl end normed_field variables [nondiscrete_normed_field 𝕜] [nondiscrete_normed_field 𝕜₂] [nondiscrete_normed_field 𝕜₃] - [normed_space 𝕜 E] [normed_space 𝕜₂ F] [normed_space 𝕜 Fₗ] + [normed_space 𝕜 E] [normed_space 𝕜 Eₗ] [normed_space 𝕜₂ F] [normed_space 𝕜 Fₗ] [normed_space 𝕜₃ G] [normed_space 𝕜 Gₗ] {σ₁₂ : 𝕜 →+* 𝕜₂} {σ₂₃ : 𝕜₂ →+* 𝕜₃} {σ₁₃ : 𝕜 →+* 𝕜₃} [ring_hom_comp_triple σ₁₂ σ₂₃ σ₁₃] @@ -307,6 +306,9 @@ begin (div_le_iff hlt).mpr $ by { apply hc })), end +theorem dist_le_op_norm (x y : E) : dist (f x) (f y) ≤ ∥f∥ * dist x y := +by simp_rw [dist_eq_norm, ← map_sub, f.le_op_norm] + theorem le_op_norm_of_le {c : ℝ} {x} (h : ∥x∥ ≤ c) : ∥f x∥ ≤ ∥f∥ * c := le_trans (f.le_op_norm x) (mul_le_mul_of_nonneg_left h f.op_norm_nonneg) @@ -444,6 +446,9 @@ instance to_normed_algebra : normed_algebra 𝕜 (E →L[𝕜] E) := theorem le_op_nnnorm : ∥f x∥₊ ≤ ∥f∥₊ * ∥x∥₊ := f.le_op_norm x +theorem nndist_le_op_nnnorm (x y : E) : nndist (f x) (f y) ≤ ∥f∥₊ * nndist x y := +dist_le_op_norm f x y + /-- continuous linear maps are Lipschitz continuous. -/ theorem lipschitz : lipschitz_with ∥f∥₊ f := (f : E →ₛₗ[σ₁₂] F).lipschitz_of_bound_nnnorm _ f.le_op_nnnorm @@ -731,6 +736,16 @@ def compL : (Fₗ →L[𝕜] Gₗ) →L[𝕜] (E →L[𝕜] Fₗ) →L[𝕜] (E @[simp] lemma compL_apply (f : Fₗ →L[𝕜] Gₗ) (g : E →L[𝕜] Fₗ) : compL 𝕜 E Fₗ Gₗ f g = f.comp g := rfl +variables (Eₗ) {𝕜 E Fₗ Gₗ} +/-- Apply `L(x,-)` pointwise to bilinear maps, as a continuous bilinear map -/ +@[simps apply] +def precompR (L : E →L[𝕜] Fₗ →L[𝕜] Gₗ) : E →L[𝕜] (Eₗ →L[𝕜] Fₗ) →L[𝕜] (Eₗ →L[𝕜] Gₗ) := +(compL 𝕜 Eₗ Fₗ Gₗ).comp L + +/-- Apply `L(-,y)` pointwise to bilinear maps, as a continuous bilinear map -/ +def precompL (L : E →L[𝕜] Fₗ →L[𝕜] Gₗ) : (Eₗ →L[𝕜] E) →L[𝕜] Fₗ →L[𝕜] (Eₗ →L[𝕜] Gₗ) := +(precompR Eₗ (flip L)).flip + section prod universes u₁ u₂ u₃ u₄ @@ -739,6 +754,7 @@ variables (M₁ : Type u₁) [semi_normed_group M₁] [normed_space 𝕜 M₁] (M₃ : Type u₃) [semi_normed_group M₃] [normed_space 𝕜 M₃] (M₄ : Type u₄) [semi_normed_group M₄] [normed_space 𝕜 M₄] +variables {Eₗ} (𝕜) /-- `continuous_linear_map.prod_map` as a continuous linear map. -/ def prod_mapL : ((M₁ →L[𝕜] M₂) × (M₃ →L[𝕜] M₄)) →L[𝕜] ((M₁ × M₃) →L[𝕜] (M₂ × M₄)) := continuous_linear_map.copy @@ -1129,7 +1145,7 @@ f.bilinear_comp (fst _ _ _) (snd _ _ _) + f.flip.bilinear_comp (snd _ _ _) (fst @[simp] lemma coe_deriv₂ (f : E →L[𝕜] Fₗ →L[𝕜] Gₗ) (p : E × Fₗ) : ⇑(f.deriv₂ p) = λ q : E × Fₗ, f p.1 q.2 + f q.1 p.2 := rfl -lemma map_add₂ (f : E →L[𝕜] Fₗ →L[𝕜] Gₗ) (x x' : E) (y y' : Fₗ) : +lemma map_add_add (f : E →L[𝕜] Fₗ →L[𝕜] Gₗ) (x x' : E) (y y' : Fₗ) : f (x + x') (y + y') = f x y + f.deriv₂ (x, y) (x', y') + f x' y' := by simp only [map_add, add_apply, coe_deriv₂, add_assoc] diff --git a/src/measure_theory/function/strongly_measurable.lean b/src/measure_theory/function/strongly_measurable.lean index e3a3b6767d8e1..3f9fc7d7c5608 100644 --- a/src/measure_theory/function/strongly_measurable.lean +++ b/src/measure_theory/function/strongly_measurable.lean @@ -1501,6 +1501,7 @@ section continuous_linear_map_nondiscrete_normed_field variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] variables {E : Type*} [normed_group E] [normed_space 𝕜 E] variables {F : Type*} [normed_group F] [normed_space 𝕜 F] +variables {G : Type*} [normed_group G] [normed_space 𝕜 G] lemma _root_.strongly_measurable.apply_continuous_linear_map {m : measurable_space α} {φ : α → F →L[𝕜] E} (hφ : strongly_measurable φ) (v : F) : @@ -1512,6 +1513,10 @@ lemma apply_continuous_linear_map {φ : α → F →L[𝕜] E} ae_strongly_measurable (λ a, φ a v) μ := (continuous_linear_map.apply 𝕜 E v).continuous.comp_ae_strongly_measurable hφ +lemma ae_strongly_measurable_comp₂ (L : E →L[𝕜] F →L[𝕜] G) {f : α → E} {g : α → F} + (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) : + ae_strongly_measurable (λ x, L (f x) (g x)) μ := +L.continuous₂.comp_ae_strongly_measurable $ hf.prod_mk hg end continuous_linear_map_nondiscrete_normed_field lemma _root_.ae_strongly_measurable_with_density_iff {E : Type*} [normed_group E] [normed_space ℝ E] From 6d2788c27600964d433697babe0861a58dd1c5aa Mon Sep 17 00:00:00 2001 From: Floris van Doorn Date: Tue, 3 May 2022 14:29:22 +0000 Subject: [PATCH 042/153] feat(analysis/calculus/cont_diff): cont_diff_succ_iff_fderiv_apply (#13797) * Prove that a map is `C^(n+1)` iff it is differentiable and all its directional derivatives in all points are `C^n`. * Also some supporting lemmas about `continuous_linear_equiv`. * We only manage to prove this when the domain is finite dimensional. * Prove one direction of `cont_diff_on_succ_iff_fderiv_within` with fewer assumptions * From the sphere eversion project Co-authored by: Patrick Massot [patrick.massot@u-psud.fr](mailto:patrick.massot@u-psud.fr) Co-authored by: Oliver Nash [github@olivernash.org](mailto:github@olivernash.org) --- src/algebra/module/equiv.lean | 5 +- src/analysis/calculus/cont_diff.lean | 100 ++++++++++++++---- .../normed_space/finite_dimension.lean | 48 ++++++++- src/analysis/normed_space/operator_norm.lean | 38 +++++++ src/topology/algebra/module/basic.lean | 2 +- 5 files changed, 169 insertions(+), 24 deletions(-) diff --git a/src/algebra/module/equiv.lean b/src/algebra/module/equiv.lean index 7a20bd3b32783..2dadf157e6a6e 100644 --- a/src/algebra/module/equiv.lean +++ b/src/algebra/module/equiv.lean @@ -227,7 +227,10 @@ include σ' omit σ' include σ₃₁ σ₂₁ σ₃₂ -@[simp] lemma symm_trans_apply +@[simp] lemma trans_symm : (e₁₂.trans e₂₃ : M₁ ≃ₛₗ[σ₁₃] M₃).symm = e₂₃.symm.trans e₁₂.symm := +rfl + +lemma symm_trans_apply (c : M₃) : (e₁₂.trans e₂₃ : M₁ ≃ₛₗ[σ₁₃] M₃).symm c = e₁₂.symm (e₂₃.symm c) := rfl omit σ₃₁ σ₂₁ σ₃₂ diff --git a/src/analysis/calculus/cont_diff.lean b/src/analysis/calculus/cont_diff.lean index fca4828fa13d1..1cd0fd7da5153 100644 --- a/src/analysis/calculus/cont_diff.lean +++ b/src/analysis/calculus/cont_diff.lean @@ -985,33 +985,38 @@ begin exact cont_diff_on_of_continuous_on_differentiable_on h.1 h.2 } end +lemma cont_diff_on_succ_of_fderiv_within {n : ℕ} (hf : differentiable_on 𝕜 f s) + (h : cont_diff_on 𝕜 n (λ y, fderiv_within 𝕜 f s y) s) : + cont_diff_on 𝕜 ((n + 1) : ℕ) f s := +begin + intros x hx, + rw [cont_diff_within_at_succ_iff_has_fderiv_within_at, insert_eq_of_mem hx], + exact ⟨s, self_mem_nhds_within, fderiv_within 𝕜 f s, + λ y hy, (hf y hy).has_fderiv_within_at, h x hx⟩ +end + /-- A function is `C^(n + 1)` on a domain with unique derivatives if and only if it is differentiable there, and its derivative (expressed with `fderiv_within`) is `C^n`. -/ theorem cont_diff_on_succ_iff_fderiv_within {n : ℕ} (hs : unique_diff_on 𝕜 s) : cont_diff_on 𝕜 ((n + 1) : ℕ) f s ↔ differentiable_on 𝕜 f s ∧ cont_diff_on 𝕜 n (λ y, fderiv_within 𝕜 f s y) s := begin - split, - { assume H, - refine ⟨H.differentiable_on (with_top.coe_le_coe.2 (nat.le_add_left 1 n)), λ x hx, _⟩, - rcases cont_diff_within_at_succ_iff_has_fderiv_within_at.1 (H x hx) - with ⟨u, hu, f', hff', hf'⟩, - rcases mem_nhds_within.1 hu with ⟨o, o_open, xo, ho⟩, - rw [inter_comm, insert_eq_of_mem hx] at ho, - have := hf'.mono ho, - rw cont_diff_within_at_inter' (mem_nhds_within_of_mem_nhds (is_open.mem_nhds o_open xo)) - at this, - apply this.congr_of_eventually_eq' _ hx, - have : o ∩ s ∈ 𝓝[s] x := mem_nhds_within.2 ⟨o, o_open, xo, subset.refl _⟩, - rw inter_comm at this, - apply filter.eventually_eq_of_mem this (λ y hy, _), - have A : fderiv_within 𝕜 f (s ∩ o) y = f' y := - ((hff' y (ho hy)).mono ho).fderiv_within (hs.inter o_open y hy), - rwa fderiv_within_inter (is_open.mem_nhds o_open hy.2) (hs y hy.1) at A, }, - { rintros ⟨hdiff, h⟩ x hx, - rw [cont_diff_within_at_succ_iff_has_fderiv_within_at, insert_eq_of_mem hx], - exact ⟨s, self_mem_nhds_within, fderiv_within 𝕜 f s, - λ y hy, (hdiff y hy).has_fderiv_within_at, h x hx⟩ } + refine ⟨λ H, _, λ h, cont_diff_on_succ_of_fderiv_within h.1 h.2⟩, + refine ⟨H.differentiable_on (with_top.coe_le_coe.2 (nat.le_add_left 1 n)), λ x hx, _⟩, + rcases cont_diff_within_at_succ_iff_has_fderiv_within_at.1 (H x hx) + with ⟨u, hu, f', hff', hf'⟩, + rcases mem_nhds_within.1 hu with ⟨o, o_open, xo, ho⟩, + rw [inter_comm, insert_eq_of_mem hx] at ho, + have := hf'.mono ho, + rw cont_diff_within_at_inter' (mem_nhds_within_of_mem_nhds (is_open.mem_nhds o_open xo)) + at this, + apply this.congr_of_eventually_eq' _ hx, + have : o ∩ s ∈ 𝓝[s] x := mem_nhds_within.2 ⟨o, o_open, xo, subset.refl _⟩, + rw inter_comm at this, + apply filter.eventually_eq_of_mem this (λ y hy, _), + have A : fderiv_within 𝕜 f (s ∩ o) y = f' y := + ((hff' y (ho hy)).mono ho).fderiv_within (hs.inter o_open y hy), + rwa fderiv_within_inter (is_open.mem_nhds o_open hy.2) (hs y hy.1) at A end /-- A function is `C^(n + 1)` on an open domain if and only if it is @@ -2734,6 +2739,59 @@ f.cont_diff_at_symm ha (hf₀'.has_fderiv_at_equiv h₀) hf end function_inverse + +/-! ### Finite dimensional results -/ +section finite_dimensional + +open function finite_dimensional +variables [complete_space 𝕜] + +/-- A family of continuous linear maps is `C^n` on `s` if all its applications are. -/ +lemma cont_diff_on_clm_apply {n : with_top ℕ} {f : E → F →L[𝕜] G} + {s : set E} [finite_dimensional 𝕜 F] : + cont_diff_on 𝕜 n f s ↔ ∀ y, cont_diff_on 𝕜 n (λ x, f x y) s := +begin + refine ⟨λ h y, (continuous_linear_map.apply 𝕜 G y).cont_diff.comp_cont_diff_on h, λ h, _⟩, + let d := finrank 𝕜 F, + have hd : d = finrank 𝕜 (fin d → 𝕜) := (finrank_fin_fun 𝕜).symm, + let e₁ := continuous_linear_equiv.of_finrank_eq hd, + let e₂ := (e₁.arrow_congr (1 : G ≃L[𝕜] G)).trans (continuous_linear_equiv.pi_ring (fin d)), + rw [← comp.left_id f, ← e₂.symm_comp_self], + exact e₂.symm.cont_diff.comp_cont_diff_on (cont_diff_on_pi.mpr (λ i, h _)) +end + +lemma cont_diff_clm_apply {n : with_top ℕ} {f : E → F →L[𝕜] G} [finite_dimensional 𝕜 F] : + cont_diff 𝕜 n f ↔ ∀ y, cont_diff 𝕜 n (λ x, f x y) := +by simp_rw [← cont_diff_on_univ, cont_diff_on_clm_apply] + +/-- This is a useful lemma to prove that a certain operation preserves functions being `C^n`. +When you do induction on `n`, this gives a useful characterization of a function being `C^(n+1)`, +assuming you have already computed the derivative. The advantage of this version over +`cont_diff_succ_iff_fderiv` is that both occurences of `cont_diff` are for functions with the same +domain and codomain (`E` and `F`). This is not the case for `cont_diff_succ_iff_fderiv`, which +often requires an inconvenient need to generalize `F`, which results in universe issues +(see the discussion in the section of `cont_diff.comp`). + +This lemma avoids these universe issues, but only applies for finite dimensional `E`. -/ +lemma cont_diff_succ_iff_fderiv_apply [finite_dimensional 𝕜 E] {n : ℕ} {f : E → F} : + cont_diff 𝕜 ((n + 1) : ℕ) f ↔ + differentiable 𝕜 f ∧ ∀ y, cont_diff 𝕜 n (λ x, fderiv 𝕜 f x y) := +by rw [cont_diff_succ_iff_fderiv, cont_diff_clm_apply] + +lemma cont_diff_on_succ_of_fderiv_apply [finite_dimensional 𝕜 E] {n : ℕ} {f : E → F} + {s : set E} (hf : differentiable_on 𝕜 f s) + (h : ∀ y, cont_diff_on 𝕜 n (λ x, fderiv_within 𝕜 f s x y) s) : + cont_diff_on 𝕜 ((n + 1) : ℕ) f s := +cont_diff_on_succ_of_fderiv_within hf $ cont_diff_on_clm_apply.mpr h + +lemma cont_diff_on_succ_iff_fderiv_apply [finite_dimensional 𝕜 E] {n : ℕ} {f : E → F} + {s : set E} (hs : unique_diff_on 𝕜 s) : cont_diff_on 𝕜 ((n + 1) : ℕ) f s ↔ + differentiable_on 𝕜 f s ∧ ∀ y, cont_diff_on 𝕜 n (λ x, fderiv_within 𝕜 f s x y) s := +by rw [cont_diff_on_succ_iff_fderiv_within hs, cont_diff_on_clm_apply] + +end finite_dimensional + + section real /-! ### Results over `ℝ` or `ℂ` diff --git a/src/analysis/normed_space/finite_dimension.lean b/src/analysis/normed_space/finite_dimension.lean index def5128210aae..6e6c7a5d82953 100644 --- a/src/analysis/normed_space/finite_dimension.lean +++ b/src/analysis/normed_space/finite_dimension.lean @@ -495,7 +495,6 @@ def basis.equiv_funL (v : basis ι 𝕜 E) : E ≃L[𝕜] (ι → 𝕜) := end, ..v.equiv_fun } - @[simp] lemma basis.constrL_apply (v : basis ι 𝕜 E) (f : ι → F) (e : E) : (v.constrL f) e = ∑ i, (v.equiv_fun e i) • f i := v.constr_apply_fintype 𝕜 _ _ @@ -739,6 +738,53 @@ begin { exact (closed_embedding_smul_left hc).is_closed_map } end +open continuous_linear_map +/-- Continuous linear equivalence between continuous linear functions `𝕜ⁿ → E` and `Eⁿ`. +The spaces `𝕜ⁿ` and `Eⁿ` are represented as `ι → 𝕜` and `ι → E`, respectively, +where `ι` is a finite type. -/ +def continuous_linear_equiv.pi_ring (ι : Type*) [fintype ι] [decidable_eq ι] : + ((ι → 𝕜) →L[𝕜] E) ≃L[𝕜] (ι → E) := +{ continuous_to_fun := + begin + refine continuous_pi (λ i, _), + exact (continuous_linear_map.apply 𝕜 E (pi.single i 1)).continuous, + end, + continuous_inv_fun := + begin + simp_rw [linear_equiv.inv_fun_eq_symm, linear_equiv.trans_symm, linear_equiv.symm_symm], + apply linear_map.continuous_of_bound _ (fintype.card ι : ℝ) (λ g, _), + rw ← nsmul_eq_mul, + apply op_norm_le_bound _ (nsmul_nonneg (norm_nonneg g) (fintype.card ι)) (λ t, _), + simp_rw [linear_map.coe_comp, linear_equiv.coe_to_linear_map, function.comp_app, + linear_map.coe_to_continuous_linear_map', linear_equiv.pi_ring_symm_apply], + apply le_trans (norm_sum_le _ _), + rw smul_mul_assoc, + refine finset.sum_le_card_nsmul _ _ _ (λ i hi, _), + rw [norm_smul, mul_comm], + exact mul_le_mul (norm_le_pi_norm g i) (norm_le_pi_norm t i) (norm_nonneg _) (norm_nonneg g), + end, + .. linear_map.to_continuous_linear_map.symm.trans (linear_equiv.pi_ring 𝕜 E ι 𝕜) } + +/-- A family of continuous linear maps is continuous on `s` if all its applications are. -/ +lemma continuous_on_clm_apply {X : Type*} [topological_space X] + [finite_dimensional 𝕜 E] {f : X → E →L[𝕜] F} {s : set X} : + continuous_on f s ↔ ∀ y, continuous_on (λ x, f x y) s := +begin + refine ⟨λ h y, (continuous_linear_map.apply 𝕜 F y).continuous.comp_continuous_on h, λ h, _⟩, + let d := finrank 𝕜 E, + have hd : d = finrank 𝕜 (fin d → 𝕜) := (finrank_fin_fun 𝕜).symm, + let e₁ : E ≃L[𝕜] fin d → 𝕜 := continuous_linear_equiv.of_finrank_eq hd, + let e₂ : (E →L[𝕜] F) ≃L[𝕜] fin d → F := + (e₁.arrow_congr (1 : F ≃L[𝕜] F)).trans (continuous_linear_equiv.pi_ring (fin d)), + rw [← function.comp.left_id f, ← e₂.symm_comp_self], + exact e₂.symm.continuous.comp_continuous_on (continuous_on_pi.mpr (λ i, h _)) +end + +lemma continuous_clm_apply {X : Type*} [topological_space X] [finite_dimensional 𝕜 E] + {f : X → E →L[𝕜] F} : + continuous f ↔ ∀ y, continuous (λ x, f x y) := +by simp_rw [continuous_iff_continuous_on_univ, continuous_on_clm_apply] + end complete_field section proper_field diff --git a/src/analysis/normed_space/operator_norm.lean b/src/analysis/normed_space/operator_norm.lean index d76cfa70fb880..08465fd5c898e 100644 --- a/src/analysis/normed_space/operator_norm.lean +++ b/src/analysis/normed_space/operator_norm.lean @@ -1743,6 +1743,44 @@ end (coord 𝕜 x h) (⟨x, submodule.mem_span_singleton_self x⟩ : 𝕜 ∙ x) = 1 := linear_equiv.coord_self 𝕜 E x h +variables {𝕜} {𝕜₄ : Type*} [nondiscrete_normed_field 𝕜₄] +variables {H : Type*} [normed_group H] [normed_space 𝕜₄ H] [normed_space 𝕜₃ G] +variables {σ₂₃ : 𝕜₂ →+* 𝕜₃} {σ₁₃ : 𝕜 →+* 𝕜₃} +variables {σ₃₄ : 𝕜₃ →+* 𝕜₄} {σ₄₃ : 𝕜₄ →+* 𝕜₃} +variables {σ₂₄ : 𝕜₂ →+* 𝕜₄} {σ₁₄ : 𝕜 →+* 𝕜₄} +variables [ring_hom_inv_pair σ₃₄ σ₄₃] [ring_hom_inv_pair σ₄₃ σ₃₄] +variables [ring_hom_comp_triple σ₂₁ σ₁₄ σ₂₄] [ring_hom_comp_triple σ₂₄ σ₄₃ σ₂₃] +variables [ring_hom_comp_triple σ₁₂ σ₂₃ σ₁₃] [ring_hom_comp_triple σ₁₃ σ₃₄ σ₁₄] +variables [ring_hom_isometric σ₁₄] [ring_hom_isometric σ₂₃] +variables [ring_hom_isometric σ₄₃] [ring_hom_isometric σ₂₄] +variables [ring_hom_isometric σ₁₃] [ring_hom_isometric σ₁₂] +variables [ring_hom_isometric σ₃₄] + +include σ₂₁ σ₃₄ σ₁₃ σ₂₄ +/-- A pair of continuous (semi)linear equivalences generates an continuous (semi)linear equivalence +between the spaces of continuous (semi)linear maps. -/ +def arrow_congrSL (e₁₂ : E ≃SL[σ₁₂] F) (e₄₃ : H ≃SL[σ₄₃] G) : + (E →SL[σ₁₄] H) ≃SL[σ₄₃] (F →SL[σ₂₃] G) := +{ map_add' := λ f g, by simp only [equiv.to_fun_as_coe, add_comp, comp_add, + continuous_linear_equiv.arrow_congr_equiv_apply], + map_smul' := λ t f, by simp only [equiv.to_fun_as_coe, smul_comp, comp_smulₛₗ, + continuous_linear_equiv.arrow_congr_equiv_apply], + continuous_to_fun := (compSL F H G σ₂₄ σ₄₃ e₄₃).continuous.comp + (continuous_linear_map.flip (compSL F E H σ₂₁ σ₁₄) e₁₂.symm).continuous, + continuous_inv_fun := (compSL E G H σ₁₃ σ₃₄ e₄₃.symm).continuous.comp + (continuous_linear_map.flip (compSL E F G σ₁₂ σ₂₃) e₁₂).continuous, + .. e₁₂.arrow_congr_equiv e₄₃, } + +omit σ₂₁ σ₃₄ σ₁₃ σ₂₄ + +/-- A pair of continuous linear equivalences generates an continuous linear equivalence between +the spaces of continuous linear maps. -/ +def arrow_congr {F H : Type*} [normed_group F] [normed_group H] + [normed_space 𝕜 F] [normed_space 𝕜 G] [normed_space 𝕜 H] + (e₁ : E ≃L[𝕜] F) (e₂ : H ≃L[𝕜] G) : + (E →L[𝕜] H) ≃L[𝕜] (F →L[𝕜] G) := +arrow_congrSL e₁ e₂ + end end continuous_linear_equiv diff --git a/src/topology/algebra/module/basic.lean b/src/topology/algebra/module/basic.lean index 2461e16ba5ee3..0be7182b4669a 100644 --- a/src/topology/algebra/module/basic.lean +++ b/src/topology/algebra/module/basic.lean @@ -1591,7 +1591,7 @@ variables {M₁} {R₄ : Type*} [semiring R₄] [module R₄ M₄] include σ₂₁ σ₃₄ σ₂₃ σ₂₄ σ₁₃ /-- A pair of continuous (semi)linear equivalences generates an equivalence between the spaces of -continuous linear maps. -/ +continuous linear maps. See also `continuous_linear_equiv.arrow_congr`. -/ @[simps] def arrow_congr_equiv (e₁₂ : M₁ ≃SL[σ₁₂] M₂) (e₄₃ : M₄ ≃SL[σ₄₃] M₃) : (M₁ →SL[σ₁₄] M₄) ≃ (M₂ →SL[σ₂₃] M₃) := { to_fun := λ f, (e₄₃ : M₄ →SL[σ₄₃] M₃).comp (f.comp (e₁₂.symm : M₂ →SL[σ₂₁] M₁)), From 36b5341d21e32968805371d91e1486aa97c88c91 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Tue, 3 May 2022 14:29:24 +0000 Subject: [PATCH 043/153] feat(ring_theory/polynomial/basic): reduce assumptions, golf (#13800) There is some reorder, so the diff is a bit large. Sorry for that. --- src/ring_theory/polynomial/basic.lean | 375 ++++++++++++++------------ 1 file changed, 205 insertions(+), 170 deletions(-) diff --git a/src/ring_theory/polynomial/basic.lean b/src/ring_theory/polynomial/basic.lean index f65af429b3f10..e698de40d457e 100644 --- a/src/ring_theory/polynomial/basic.lean +++ b/src/ring_theory/polynomial/basic.lean @@ -29,13 +29,17 @@ noncomputable theory open_locale classical big_operators polynomial universes u v w +variables {R : Type u} {S : Type*} namespace polynomial -instance {R : Type u} [semiring R] (p : ℕ) [h : char_p R p] : char_p R[X] p := +section semiring +variables [semiring R] + +instance (p : ℕ) [h : char_p R p] : char_p R[X] p := let ⟨h⟩ := h in ⟨λ n, by rw [← map_nat_cast C, ← C_0, C_inj, h]⟩ -variables (R : Type u) [comm_ring R] +variables (R) /-- The `R`-submodule of `R[X]` consisting of polynomials of degree ≤ `n`. -/ def degree_le (n : with_bot ℕ) : submodule R R[X] := @@ -101,7 +105,7 @@ begin end /-- The first `n` coefficients on `degree_lt n` form a linear equivalence with `fin n → R`. -/ -def degree_lt_equiv (R : Type u) [comm_ring R] (n : ℕ) : degree_lt R n ≃ₗ[R] (fin n → R) := +def degree_lt_equiv (n : ℕ) : degree_lt R n ≃ₗ[R] (fin n → R) := { to_fun := λ p n, (↑p : R[X]).coeff n, inv_fun := λ f, ⟨∑ i : fin n, monomial i (f i), (degree_lt R n).sum_mem (λ i _, mem_degree_lt.mpr (lt_of_le_of_lt @@ -153,12 +157,63 @@ begin exact ⟨n, h, rfl⟩, end +lemma geom_sum_X_comp_X_add_one_eq_sum (n : ℕ) : + (geom_sum (X : R[X]) n).comp (X + 1) = + (finset.range n).sum (λ (i : ℕ), (n.choose (i + 1) : R[X]) * X ^ i) := +begin + ext i, + transitivity (n.choose (i + 1) : R), swap, + { simp only [finset_sum_coeff, ← C_eq_nat_cast, coeff_C_mul_X_pow], + rw [finset.sum_eq_single i, if_pos rfl], + { simp only [@eq_comm _ i, if_false, eq_self_iff_true, implies_true_iff] {contextual := tt}, }, + { simp only [nat.lt_add_one_iff, nat.choose_eq_zero_of_lt, nat.cast_zero, finset.mem_range, + not_lt, eq_self_iff_true, if_true, implies_true_iff] {contextual := tt}, } }, + induction n with n ih generalizing i, + { simp only [geom_sum_zero, zero_comp, coeff_zero, nat.choose_zero_succ, nat.cast_zero], }, + simp only [geom_sum_succ', ih, add_comp, X_pow_comp, coeff_add, nat.choose_succ_succ, + nat.cast_add, coeff_X_add_one_pow], +end + +lemma monic.geom_sum {P : R[X]} + (hP : P.monic) (hdeg : 0 < P.nat_degree) {n : ℕ} (hn : n ≠ 0) : (geom_sum P n).monic := +begin + nontriviality R, + cases n, { exact (hn rfl).elim }, + rw [geom_sum_succ', geom_sum_def], + refine (hP.pow _).add_of_left _, + refine lt_of_le_of_lt (degree_sum_le _ _) _, + rw [finset.sup_lt_iff], + { simp only [finset.mem_range, degree_eq_nat_degree (hP.pow _).ne_zero, + with_bot.coe_lt_coe, hP.nat_degree_pow], + intro k, exact nsmul_lt_nsmul hdeg }, + { rw [bot_lt_iff_ne_bot, ne.def, degree_eq_bot], + exact (hP.pow _).ne_zero } +end + +lemma monic.geom_sum' {P : R[X]} + (hP : P.monic) (hdeg : 0 < P.degree) {n : ℕ} (hn : n ≠ 0) : (geom_sum P n).monic := +hP.geom_sum (nat_degree_pos_iff_degree_pos.2 hdeg) hn + +lemma monic_geom_sum_X {n : ℕ} (hn : n ≠ 0) : + (geom_sum (X : R[X]) n).monic := +begin + nontriviality R, + apply monic_X.geom_sum _ hn, + simpa only [nat_degree_X] using zero_lt_one +end + +end semiring + +section ring +variables [ring R] + /-- Given a polynomial, return the polynomial whose coefficients are in the ring closure of the original coefficients. -/ def restriction (p : R[X]) : polynomial (subring.closure (↑p.frange : set R)) := ∑ i in p.support, monomial i (⟨p.coeff i, if H : p.coeff i = 0 then H.symm ▸ (subring.closure _).zero_mem - else subring.subset_closure (p.coeff_mem_frange _ H)⟩ : (subring.closure (↑p.frange : set R))) + else subring.subset_closure (p.coeff_mem_frange _ H)⟩ : + (subring.closure (↑p.frange : set R))) @[simp] theorem coeff_restriction {p : R[X]} {n : ℕ} : ↑(coeff (restriction p) n) = coeff p n := @@ -183,7 +238,8 @@ begin exact ⟨λ H, by { rw H, refl }, λ H, subtype.coe_injective H⟩ end -@[simp] theorem map_restriction (p : R[X]) : p.restriction.map (algebra_map _ _) = p := +@[simp] theorem map_restriction {R : Type u} [comm_ring R] + (p : R[X]) : p.restriction.map (algebra_map _ _) = p := ext $ λ n, by rw [coeff_map, algebra.algebra_map_of_subring_apply, coeff_restriction] @[simp] theorem degree_restriction {p : R[X]} : (restriction p).degree = p.degree := @@ -206,7 +262,7 @@ by simp only [restriction, finset.sum_empty, support_zero] @[simp] theorem restriction_one : restriction (1 : R[X]) = 1 := ext $ λ i, subtype.eq $ by rw [coeff_restriction', coeff_one, coeff_one]; split_ifs; refl -variables {S : Type v} [ring S] {f : R →+* S} {x : S} +variables [semiring S] {f : R →+* S} {x : S} theorem eval₂_restriction {p : R[X]} : eval₂ f x p = @@ -216,56 +272,6 @@ begin refl, end -lemma geom_sum_X_comp_X_add_one_eq_sum (n : ℕ) : - (geom_sum (X : R[X]) n).comp (X + 1) = - (finset.range n).sum (λ (i : ℕ), (n.choose (i + 1) : R[X]) * X ^ i) := -begin - ext i, - transitivity (n.choose (i + 1) : R), swap, - { simp only [finset_sum_coeff, ← C_eq_nat_cast, coeff_C_mul_X_pow], - rw [finset.sum_eq_single i, if_pos rfl], - { simp only [@eq_comm _ i, if_false, eq_self_iff_true, implies_true_iff] {contextual := tt}, }, - { simp only [nat.lt_add_one_iff, nat.choose_eq_zero_of_lt, nat.cast_zero, finset.mem_range, - not_lt, eq_self_iff_true, if_true, implies_true_iff] {contextual := tt}, } }, - induction n with n ih generalizing i, - { simp only [geom_sum_zero, zero_comp, coeff_zero, nat.choose_zero_succ, nat.cast_zero], }, - simp only [geom_sum_succ', ih, add_comp, pow_comp, X_comp, coeff_add, nat.choose_succ_succ, - nat.cast_add, add_pow, one_pow, mul_one, finset_sum_coeff, ← C_eq_nat_cast, mul_comm _ (C _), - coeff_C_mul_X_pow], - rw [finset.sum_eq_single i, if_pos rfl], - { simp only [@eq_comm _ i, if_false, eq_self_iff_true, implies_true_iff] {contextual := tt}, }, - { simp only [nat.lt_add_one_iff, nat.choose_eq_zero_of_lt, nat.cast_zero, finset.mem_range, - eq_self_iff_true, if_true, implies_true_iff, not_le] {contextual := tt}, }, -end - -lemma monic.geom_sum {R : Type*} [semiring R] {P : R[X]} - (hP : P.monic) (hdeg : 0 < P.nat_degree) {n : ℕ} (hn : n ≠ 0) : (geom_sum P n).monic := -begin - nontriviality R, - cases n, { exact (hn rfl).elim }, - rw [geom_sum_succ', geom_sum_def], - refine (hP.pow _).add_of_left _, - refine lt_of_le_of_lt (degree_sum_le _ _) _, - rw [finset.sup_lt_iff], - { simp only [finset.mem_range, degree_eq_nat_degree (hP.pow _).ne_zero, - with_bot.coe_lt_coe, hP.nat_degree_pow], - intro k, exact nsmul_lt_nsmul hdeg }, - { rw [bot_lt_iff_ne_bot, ne.def, degree_eq_bot], - exact (hP.pow _).ne_zero } -end - -lemma monic.geom_sum' {R : Type*} [semiring R] {P : R[X]} - (hP : P.monic) (hdeg : 0 < P.degree) {n : ℕ} (hn : n ≠ 0) : (geom_sum P n).monic := -hP.geom_sum (nat_degree_pos_iff_degree_pos.2 hdeg) hn - -lemma monic_geom_sum_X (R : Type*) [semiring R] {n : ℕ} (hn : n ≠ 0) : - (geom_sum (X : R[X]) n).monic := -begin - nontriviality R, - apply monic_X.geom_sum _ hn, - simpa only [nat_degree_X] using zero_lt_one -end - section to_subring variables (p : R[X]) (T : subring R) @@ -355,6 +361,11 @@ begin exact subtype.mem (coeff p n : T) end +end ring + +section comm_ring +variables [comm_ring R] + section mod_by_monic variables {q : R[X]} @@ -369,14 +380,47 @@ submodule.ext (λ f, (mem_ker_mod_by_monic hq).trans ideal.mem_span_singleton.sy end mod_by_monic -end polynomial +end comm_ring -variables {R : Type u} {S : Type*} {σ : Type v} {M : Type w} -variables [comm_ring R] [comm_ring S] [add_comm_group M] [module R M] +end polynomial namespace ideal open polynomial +section semiring +variables [semiring R] + +/-- Transport an ideal of `R[X]` to an `R`-submodule of `R[X]`. -/ +def of_polynomial (I : ideal R[X]) : submodule R R[X] := +{ carrier := I.carrier, + zero_mem' := I.zero_mem, + add_mem' := λ _ _, I.add_mem, + smul_mem' := λ c x H, by { rw [← C_mul'], exact I.mul_mem_left _ H } } + +variables {I : ideal R[X]} +theorem mem_of_polynomial (x) : x ∈ I.of_polynomial ↔ x ∈ I := iff.rfl +variables (I) + +/-- Given an ideal `I` of `R[X]`, make the `R`-submodule of `I` +consisting of polynomials of degree ≤ `n`. -/ +def degree_le (n : with_bot ℕ) : submodule R R[X] := +degree_le R n ⊓ I.of_polynomial + +/-- Given an ideal `I` of `R[X]`, make the ideal in `R` of +leading coefficients of polynomials in `I` with degree ≤ `n`. -/ +def leading_coeff_nth (n : ℕ) : ideal R := +(I.degree_le n).map $ lcoeff R n + +/-- Given an ideal `I` in `R[X]`, make the ideal in `R` of the +leading coefficients in `I`. -/ +def leading_coeff : ideal R := +⨆ n : ℕ, I.leading_coeff_nth n + +end semiring + +section comm_semiring +variables [comm_semiring R] [semiring S] + /-- If every coefficient of a polynomial is in an ideal `I`, then so is the polynomial itself -/ lemma polynomial_mem_ideal_of_coeff_mem_ideal (I : ideal R[X]) (p : R[X]) (hp : ∀ (n : ℕ), (p.coeff n) ∈ I.comap C) : p ∈ I := @@ -417,6 +461,102 @@ begin simp_rw [coe_map_ring_hom, coeff_map, coeff_zero, ring_hom.mem_ker], end +variable (I : ideal R[X]) + +theorem mem_leading_coeff_nth (n : ℕ) (x) : + x ∈ I.leading_coeff_nth n ↔ ∃ p ∈ I, degree p ≤ n ∧ p.leading_coeff = x := +begin + simp only [leading_coeff_nth, degree_le, submodule.mem_map, lcoeff_apply, submodule.mem_inf, + mem_degree_le], + split, + { rintro ⟨p, ⟨hpdeg, hpI⟩, rfl⟩, + cases lt_or_eq_of_le hpdeg with hpdeg hpdeg, + { refine ⟨0, I.zero_mem, bot_le, _⟩, + rw [leading_coeff_zero, eq_comm], + exact coeff_eq_zero_of_degree_lt hpdeg }, + { refine ⟨p, hpI, le_of_eq hpdeg, _⟩, + rw [polynomial.leading_coeff, nat_degree, hpdeg], refl } }, + { rintro ⟨p, hpI, hpdeg, rfl⟩, + have : nat_degree p + (n - nat_degree p) = n, + { exact add_tsub_cancel_of_le (nat_degree_le_of_degree_le hpdeg) }, + refine ⟨p * X ^ (n - nat_degree p), ⟨_, I.mul_mem_right _ hpI⟩, _⟩, + { apply le_trans (degree_mul_le _ _) _, + apply le_trans (add_le_add (degree_le_nat_degree) (degree_X_pow_le _)) _, + rw [← with_bot.coe_add, this], + exact le_rfl }, + { rw [polynomial.leading_coeff, ← coeff_mul_X_pow p (n - nat_degree p), this] } } +end + +theorem mem_leading_coeff_nth_zero (x) : + x ∈ I.leading_coeff_nth 0 ↔ C x ∈ I := +(mem_leading_coeff_nth _ _ _).trans +⟨λ ⟨p, hpI, hpdeg, hpx⟩, by rwa [← hpx, polynomial.leading_coeff, + nat.eq_zero_of_le_zero (nat_degree_le_of_degree_le hpdeg), + ← eq_C_of_degree_le_zero hpdeg], +λ hx, ⟨C x, hx, degree_C_le, leading_coeff_C x⟩⟩ + +theorem leading_coeff_nth_mono {m n : ℕ} (H : m ≤ n) : + I.leading_coeff_nth m ≤ I.leading_coeff_nth n := +begin + intros r hr, + simp only [set_like.mem_coe, mem_leading_coeff_nth] at hr ⊢, + rcases hr with ⟨p, hpI, hpdeg, rfl⟩, + refine ⟨p * X ^ (n - m), I.mul_mem_right _ hpI, _, leading_coeff_mul_X_pow⟩, + refine le_trans (degree_mul_le _ _) _, + refine le_trans (add_le_add hpdeg (degree_X_pow_le _)) _, + rw [← with_bot.coe_add, add_tsub_cancel_of_le H], + exact le_rfl +end + +theorem mem_leading_coeff (x) : + x ∈ I.leading_coeff ↔ ∃ p ∈ I, polynomial.leading_coeff p = x := +begin + rw [leading_coeff, submodule.mem_supr_of_directed], + simp only [mem_leading_coeff_nth], + { split, { rintro ⟨i, p, hpI, hpdeg, rfl⟩, exact ⟨p, hpI, rfl⟩ }, + rintro ⟨p, hpI, rfl⟩, exact ⟨nat_degree p, p, hpI, degree_le_nat_degree, rfl⟩ }, + intros i j, exact ⟨i + j, I.leading_coeff_nth_mono (nat.le_add_right _ _), + I.leading_coeff_nth_mono (nat.le_add_left _ _)⟩ +end + +end comm_semiring + +section ring +variables [ring R] + +/-- `polynomial R` is never a field for any ring `R`. -/ +lemma polynomial_not_is_field : ¬ is_field R[X] := +begin + by_contradiction hR, + by_cases hR' : ∃ (x y : R), x ≠ y, + { haveI : nontrivial R := let ⟨x, y, hxy⟩ := hR' in nontrivial_of_ne x y hxy, + obtain ⟨p, hp⟩ := hR.mul_inv_cancel X_ne_zero, + by_cases hp0 : p = 0, + { replace hp := congr_arg degree hp, + rw [hp0, mul_zero, degree_zero, degree_one] at hp, + contradiction }, + { have : p.degree < (X * p).degree := (X_mul.symm : p * X = _) ▸ degree_lt_degree_mul_X hp0, + rw [congr_arg degree hp, degree_one, nat.with_bot.lt_zero_iff, degree_eq_bot] at this, + exact hp0 this } }, + { push_neg at hR', + exact let ⟨x, y, hxy⟩ := hR.exists_pair_ne in hxy (polynomial.ext (λ n, hR' _ _)) } +end + +/-- The only constant in a maximal ideal over a field is `0`. -/ +lemma eq_zero_of_constant_mem_of_maximal (hR : is_field R) + (I : ideal R[X]) [hI : I.is_maximal] (x : R) (hx : C x ∈ I) : x = 0 := +begin + refine classical.by_contradiction (λ hx0, hI.ne_top ((eq_top_iff_one I).2 _)), + obtain ⟨y, hy⟩ := hR.mul_inv_cancel hx0, + convert I.mul_mem_left (C y) hx, + rw [← C.map_mul, hR.mul_comm y x, hy, ring_hom.map_one], +end + +end ring + +section comm_ring +variables [comm_ring R] + lemma quotient_map_C_eq_zero {I : ideal R} : ∀ a ∈ I, ((quotient.mk (map C I : ideal R[X])).comp C) a = 0 := begin @@ -536,123 +676,18 @@ begin exact hx, end -/-- `polynomial R` is never a field for any ring `R`. -/ -lemma polynomial_not_is_field : ¬ is_field R[X] := -begin - by_contradiction hR, - by_cases hR' : ∃ (x y : R), x ≠ y, - { haveI : nontrivial R := let ⟨x, y, hxy⟩ := hR' in nontrivial_of_ne x y hxy, - obtain ⟨p, hp⟩ := hR.mul_inv_cancel X_ne_zero, - by_cases hp0 : p = 0, - { replace hp := congr_arg degree hp, - rw [hp0, mul_zero, degree_zero, degree_one] at hp, - contradiction }, - { have : p.degree < (X * p).degree := (mul_comm p X) ▸ degree_lt_degree_mul_X hp0, - rw [congr_arg degree hp, degree_one, nat.with_bot.lt_zero_iff, degree_eq_bot] at this, - exact hp0 this } }, - { push_neg at hR', - exact let ⟨x, y, hxy⟩ := hR.exists_pair_ne in hxy (polynomial.ext (λ n, hR' _ _)) } -end - -/-- The only constant in a maximal ideal over a field is `0`. -/ -lemma eq_zero_of_constant_mem_of_maximal (hR : is_field R) - (I : ideal R[X]) [hI : I.is_maximal] (x : R) (hx : C x ∈ I) : x = 0 := -begin - refine classical.by_contradiction (λ hx0, hI.ne_top ((eq_top_iff_one I).2 _)), - obtain ⟨y, hy⟩ := hR.mul_inv_cancel hx0, - convert I.mul_mem_left (C y) hx, - rw [← C.map_mul, mul_comm y x, hy, ring_hom.map_one], -end - -/-- Transport an ideal of `R[X]` to an `R`-submodule of `R[X]`. -/ -def of_polynomial (I : ideal R[X]) : submodule R R[X] := -{ carrier := I.carrier, - zero_mem' := I.zero_mem, - add_mem' := λ _ _, I.add_mem, - smul_mem' := λ c x H, by { rw [← C_mul'], exact I.mul_mem_left _ H } } - -variables {I : ideal R[X]} -theorem mem_of_polynomial (x) : x ∈ I.of_polynomial ↔ x ∈ I := iff.rfl -variables (I) - -/-- Given an ideal `I` of `R[X]`, make the `R`-submodule of `I` -consisting of polynomials of degree ≤ `n`. -/ -def degree_le (n : with_bot ℕ) : submodule R R[X] := -degree_le R n ⊓ I.of_polynomial - -/-- Given an ideal `I` of `R[X]`, make the ideal in `R` of -leading coefficients of polynomials in `I` with degree ≤ `n`. -/ -def leading_coeff_nth (n : ℕ) : ideal R := -(I.degree_le n).map $ lcoeff R n - -theorem mem_leading_coeff_nth (n : ℕ) (x) : - x ∈ I.leading_coeff_nth n ↔ ∃ p ∈ I, degree p ≤ n ∧ leading_coeff p = x := -begin - simp only [leading_coeff_nth, degree_le, submodule.mem_map, lcoeff_apply, submodule.mem_inf, - mem_degree_le], - split, - { rintro ⟨p, ⟨hpdeg, hpI⟩, rfl⟩, - cases lt_or_eq_of_le hpdeg with hpdeg hpdeg, - { refine ⟨0, I.zero_mem, bot_le, _⟩, - rw [leading_coeff_zero, eq_comm], - exact coeff_eq_zero_of_degree_lt hpdeg }, - { refine ⟨p, hpI, le_of_eq hpdeg, _⟩, - rw [leading_coeff, nat_degree, hpdeg], refl } }, - { rintro ⟨p, hpI, hpdeg, rfl⟩, - have : nat_degree p + (n - nat_degree p) = n, - { exact add_tsub_cancel_of_le (nat_degree_le_of_degree_le hpdeg) }, - refine ⟨p * X ^ (n - nat_degree p), ⟨_, I.mul_mem_right _ hpI⟩, _⟩, - { apply le_trans (degree_mul_le _ _) _, - apply le_trans (add_le_add (degree_le_nat_degree) (degree_X_pow_le _)) _, - rw [← with_bot.coe_add, this], - exact le_rfl }, - { rw [leading_coeff, ← coeff_mul_X_pow p (n - nat_degree p), this] } } -end - -theorem mem_leading_coeff_nth_zero (x) : - x ∈ I.leading_coeff_nth 0 ↔ C x ∈ I := -(mem_leading_coeff_nth _ _ _).trans -⟨λ ⟨p, hpI, hpdeg, hpx⟩, by rwa [← hpx, leading_coeff, - nat.eq_zero_of_le_zero (nat_degree_le_of_degree_le hpdeg), - ← eq_C_of_degree_le_zero hpdeg], -λ hx, ⟨C x, hx, degree_C_le, leading_coeff_C x⟩⟩ - -theorem leading_coeff_nth_mono {m n : ℕ} (H : m ≤ n) : - I.leading_coeff_nth m ≤ I.leading_coeff_nth n := -begin - intros r hr, - simp only [set_like.mem_coe, mem_leading_coeff_nth] at hr ⊢, - rcases hr with ⟨p, hpI, hpdeg, rfl⟩, - refine ⟨p * X ^ (n - m), I.mul_mem_right _ hpI, _, leading_coeff_mul_X_pow⟩, - refine le_trans (degree_mul_le _ _) _, - refine le_trans (add_le_add hpdeg (degree_X_pow_le _)) _, - rw [← with_bot.coe_add, add_tsub_cancel_of_le H], - exact le_rfl -end - -/-- Given an ideal `I` in `R[X]`, make the ideal in `R` of the -leading coefficients in `I`. -/ -def leading_coeff : ideal R := -⨆ n : ℕ, I.leading_coeff_nth n - -theorem mem_leading_coeff (x) : - x ∈ I.leading_coeff ↔ ∃ p ∈ I, polynomial.leading_coeff p = x := -begin - rw [leading_coeff, submodule.mem_supr_of_directed], - simp only [mem_leading_coeff_nth], - { split, { rintro ⟨i, p, hpI, hpdeg, rfl⟩, exact ⟨p, hpI, rfl⟩ }, - rintro ⟨p, hpI, rfl⟩, exact ⟨nat_degree p, p, hpI, degree_le_nat_degree, rfl⟩ }, - intros i j, exact ⟨i + j, I.leading_coeff_nth_mono (nat.le_add_right _ _), - I.leading_coeff_nth_mono (nat.le_add_left _ _)⟩ -end - -theorem is_fg_degree_le [is_noetherian_ring R] (n : ℕ) : +theorem is_fg_degree_le [is_noetherian_ring R] (I : ideal R[X]) (n : ℕ) : submodule.fg (I.degree_le n) := is_noetherian_submodule_left.1 (is_noetherian_of_fg_of_noetherian _ ⟨_, degree_le_eq_span_X_pow.symm⟩) _ +end comm_ring + end ideal +variables {σ : Type v} {M : Type w} +variables [comm_ring R] [comm_ring S] [add_comm_group M] [module R M] + section prime variables (σ) {r : R} From 2f38ccb4d5529facc13e13b7d049199ac2db52e9 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 3 May 2022 14:29:24 +0000 Subject: [PATCH 044/153] chore(data/matrix/basic): add lemmas about powers of matrices (#13815) Shows that: * natural powers commute with `transpose`, `conj_transpose`, `diagonal`, `block_diagonal`, and `block_diagonal'`. * integer powers commute with `transpose`, and `conj_transpose`. --- src/data/matrix/basic.lean | 32 +++++++++++++++++++++++++---- src/data/matrix/block.lean | 10 +++++++++ src/linear_algebra/matrix/zpow.lean | 10 +++++++++ 3 files changed, 48 insertions(+), 4 deletions(-) diff --git a/src/data/matrix/basic.lean b/src/data/matrix/basic.lean index 443f84a8efc68..a3dd7ba1f9589 100644 --- a/src/data/matrix/basic.lean +++ b/src/data/matrix/basic.lean @@ -84,6 +84,10 @@ lemma map_map {M : matrix m n α} {β γ : Type*} {f : α → β} {g : β → γ (M.map f).map g = M.map (g ∘ f) := by { ext, refl, } +lemma map_injective {f : α → β} (hf : function.injective f) : + function.injective (λ M : matrix m n α, M.map f) := +λ M N h, ext $ λ i j, hf $ ext_iff.mpr h i j + /-- The transpose of a matrix. -/ def transpose (M : matrix m n α) : matrix n m α | x y := M y x @@ -714,6 +718,10 @@ instance [fintype n] [decidable_eq n] [ring α] : ring (matrix n n α) := section semiring variables [semiring α] +lemma diagonal_pow [fintype n] [decidable_eq n] (v : n → α) (k : ℕ) : + diagonal v ^ k = diagonal (v ^ k) := +(map_pow (diagonal_ring_hom n α) v k).symm + @[simp] lemma mul_mul_left [fintype n] (M : matrix m n α) (N : matrix n o α) (a : α) : (λ i j, a * M i j) ⬝ N = a • (M ⬝ N) := smul_mul a M N @@ -1271,6 +1279,8 @@ lemma transpose_sum [add_comm_monoid α] {ι : Type*} (s : finset ι) (M : ι (∑ i in s, M i)ᵀ = ∑ i in s, (M i)ᵀ := (transpose_add_equiv : matrix m n α ≃+ matrix n m α).to_add_monoid_hom.map_sum _ s +variables (m α) + /-- `matrix.transpose` as a `ring_equiv` to the opposite ring -/ @[simps] def transpose_ring_equiv [add_comm_monoid α] [comm_semigroup α] [fintype m] : @@ -1281,9 +1291,15 @@ def transpose_ring_equiv [add_comm_monoid α] [comm_semigroup α] [fintype m] : (mul_opposite.op_mul _ _), ..transpose_add_equiv.trans mul_opposite.op_add_equiv } +variables {m α} + +@[simp] lemma transpose_pow [comm_semiring α] [fintype m] [decidable_eq m] (M : matrix m m α) + (k : ℕ) : (M ^ k)ᵀ = Mᵀ ^ k := +mul_opposite.op_injective $ map_pow (transpose_ring_equiv m α) M k + lemma transpose_list_prod [comm_semiring α] [fintype m] [decidable_eq m] (l : list (matrix m m α)) : l.prodᵀ = (l.map transpose).reverse.prod := -(transpose_ring_equiv : matrix m m α ≃+* (matrix m m α)ᵐᵒᵖ).unop_map_list_prod l +(transpose_ring_equiv m α).unop_map_list_prod l end transpose @@ -1356,9 +1372,11 @@ lemma conj_transpose_sum [add_comm_monoid α] [star_add_monoid α] {ι : Type*} (∑ i in s, M i)ᴴ = ∑ i in s, (M i)ᴴ := (conj_transpose_add_equiv : matrix m n α ≃+ matrix n m α).to_add_monoid_hom.map_sum _ s +variables (m α) + /-- `matrix.conj_transpose` as a `ring_equiv` to the opposite ring -/ @[simps] -def conj_transpose_ring_equiv [comm_semiring α] [star_ring α] [fintype m] : +def conj_transpose_ring_equiv [semiring α] [star_ring α] [fintype m] : matrix m m α ≃+* (matrix m m α)ᵐᵒᵖ := { to_fun := λ M, mul_opposite.op (Mᴴ), inv_fun := λ M, M.unopᴴ, @@ -1366,10 +1384,16 @@ def conj_transpose_ring_equiv [comm_semiring α] [star_ring α] [fintype m] : (mul_opposite.op_mul _ _), ..conj_transpose_add_equiv.trans mul_opposite.op_add_equiv } -lemma conj_transpose_list_prod [comm_semiring α] [star_ring α] [fintype m] [decidable_eq m] +variables {m α} + +@[simp] lemma conj_transpose_pow [semiring α] [star_ring α] [fintype m] [decidable_eq m] + (M : matrix m m α) (k : ℕ) : (M ^ k)ᴴ = Mᴴ ^ k := +mul_opposite.op_injective $ map_pow (conj_transpose_ring_equiv m α) M k + +lemma conj_transpose_list_prod [semiring α] [star_ring α] [fintype m] [decidable_eq m] (l : list (matrix m m α)) : l.prodᴴ = (l.map conj_transpose).reverse.prod := -(conj_transpose_ring_equiv : matrix m m α ≃+* (matrix m m α)ᵐᵒᵖ).unop_map_list_prod l +(conj_transpose_ring_equiv m α).unop_map_list_prod l end conj_transpose diff --git a/src/data/matrix/block.lean b/src/data/matrix/block.lean index f88cae4d27cab..cf1c1a74e5137 100644 --- a/src/data/matrix/block.lean +++ b/src/data/matrix/block.lean @@ -323,6 +323,11 @@ def block_diagonal_ring_hom [decidable_eq m] [fintype o] [fintype m] [non_assoc_ ..block_diagonal_add_monoid_hom m m o α } end +@[simp] lemma block_diagonal_pow [decidable_eq m] [fintype o] [fintype m] [semiring α] + (M : o → matrix m m α) (n : ℕ) : + block_diagonal (M ^ n) = block_diagonal M ^ n := +map_pow (block_diagonal_ring_hom m o α) M n + @[simp] lemma block_diagonal_smul {R : Type*} [monoid R] [add_monoid α] [distrib_mul_action R α] (x : R) (M : o → matrix m n α) : block_diagonal (x • M) = x • block_diagonal M := by { ext, simp only [block_diagonal_apply, pi.smul_apply], split_ifs; simp } @@ -465,6 +470,11 @@ def block_diagonal'_ring_hom [Π i, decidable_eq (m' i)] [fintype o] [Π i, fint ..block_diagonal'_add_monoid_hom m' m' α } end +@[simp] lemma block_diagonal'_pow [Π i, decidable_eq (m' i)] [fintype o] [Π i, fintype (m' i)] + [semiring α] (M : Π i, matrix (m' i) (m' i) α) (n : ℕ) : + block_diagonal' (M ^ n) = block_diagonal' M ^ n := +map_pow (block_diagonal'_ring_hom m' α) M n + @[simp] lemma block_diagonal'_smul {R : Type*} [semiring R] [add_comm_monoid α] [module R α] (x : R) (M : Π i, matrix (m' i) (n' i) α) : block_diagonal' (x • M) = x • block_diagonal' M := by { ext, simp only [block_diagonal'_apply, pi.smul_apply], split_ifs; simp } diff --git a/src/linear_algebra/matrix/zpow.lean b/src/linear_algebra/matrix/zpow.lean index 4705f34edc653..8e31e3d706ec6 100644 --- a/src/linear_algebra/matrix/zpow.lean +++ b/src/linear_algebra/matrix/zpow.lean @@ -312,6 +312,16 @@ theorem one_div_zpow {A : M} (n : ℤ) : (1 / A) ^ n = 1 / A ^ n := by simp only [one_div, inv_zpow] +@[simp] theorem transpose_zpow (A : M) : ∀ (n : ℤ), (A ^ n)ᵀ = Aᵀ ^ n +| (n : ℕ) := by rw [zpow_coe_nat, zpow_coe_nat, transpose_pow] +| -[1+ n] := by + rw [zpow_neg_succ_of_nat, zpow_neg_succ_of_nat, transpose_nonsing_inv, transpose_pow] + +@[simp] theorem conj_transpose_zpow [star_ring R] (A : M) : ∀ (n : ℤ), (A ^ n)ᴴ = Aᴴ ^ n +| (n : ℕ) := by rw [zpow_coe_nat, zpow_coe_nat, conj_transpose_pow] +| -[1+ n] := by + rw [zpow_neg_succ_of_nat, zpow_neg_succ_of_nat, conj_transpose_nonsing_inv, conj_transpose_pow] + end zpow end matrix From 8a5b4a7c1e59fc8bf83d74ead11a9bfaee1b0a52 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 3 May 2022 14:29:25 +0000 Subject: [PATCH 045/153] =?UTF-8?q?feat(analysis/special=5Ffunctions/compl?= =?UTF-8?q?ex/arg):=20lemmas=20about=20`arg=20z`=20and=20`=C2=B1(=CF=80=20?= =?UTF-8?q?/=202)`=20(#13821)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- .../special_functions/complex/arg.lean | 35 +++++++++++++++++++ 1 file changed, 35 insertions(+) diff --git a/src/analysis/special_functions/complex/arg.lean b/src/analysis/special_functions/complex/arg.lean index 2b1de2b67080f..6c1e42e0fa2aa 100644 --- a/src/analysis/special_functions/complex/arg.lean +++ b/src/analysis/special_functions/complex/arg.lean @@ -278,6 +278,41 @@ begin { exact arg_real_mul (conj x) (by simp [hx]) } end +lemma arg_le_pi_div_two_iff {z : ℂ} : arg z ≤ π / 2 ↔ 0 ≤ re z ∨ im z < 0 := +begin + cases le_or_lt 0 (re z) with hre hre, + { simp only [hre, arg_of_re_nonneg hre, real.arcsin_le_pi_div_two, true_or] }, + simp only [hre.not_le, false_or], + cases le_or_lt 0 (im z) with him him, + { simp only [him.not_lt], + rw [iff_false, not_le, arg_of_re_neg_of_im_nonneg hre him, ← sub_lt_iff_lt_add, half_sub, + real.neg_pi_div_two_lt_arcsin, neg_im, neg_div, neg_lt_neg_iff, div_lt_one, ← + _root_.abs_of_nonneg him, abs_im_lt_abs], + exacts [hre.ne, abs_pos.2 $ ne_of_apply_ne re hre.ne] }, + { simp only [him], + rw [iff_true, arg_of_re_neg_of_im_neg hre him], + exact (sub_le_self _ real.pi_pos.le).trans (real.arcsin_le_pi_div_two _) } +end + +lemma neg_pi_div_two_le_arg_iff {z : ℂ} : -(π / 2) ≤ arg z ↔ 0 ≤ re z ∨ 0 ≤ im z := +begin + cases le_or_lt 0 (re z) with hre hre, + { simp only [hre, arg_of_re_nonneg hre, real.neg_pi_div_two_le_arcsin, true_or] }, + simp only [hre.not_le, false_or], + cases le_or_lt 0 (im z) with him him, + { simp only [him], + rw [iff_true, arg_of_re_neg_of_im_nonneg hre him], + exact (real.neg_pi_div_two_le_arcsin _).trans (le_add_of_nonneg_right real.pi_pos.le) }, + { simp only [him.not_le], + rw [iff_false, not_le, arg_of_re_neg_of_im_neg hre him, sub_lt_iff_lt_add', ← sub_eq_add_neg, + sub_half, real.arcsin_lt_pi_div_two, div_lt_one, neg_im, ← abs_of_neg him, abs_im_lt_abs], + exacts [hre.ne, abs_pos.2 $ ne_of_apply_ne re hre.ne] } +end + +@[simp] lemma abs_arg_le_pi_div_two_iff {z : ℂ} : |arg z| ≤ π / 2 ↔ 0 ≤ re z := +by rw [abs_le, arg_le_pi_div_two_iff, neg_pi_div_two_le_arg_iff, ← or_and_distrib_left, ← not_le, + and_not_self, or_false] + @[simp] lemma arg_conj_coe_angle (x : ℂ) : (arg (conj x) : real.angle) = -arg x := begin by_cases h : arg x = π; From 4cea0a8fc0d71a8c906faa887890c62ba52495a0 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 3 May 2022 14:29:27 +0000 Subject: [PATCH 046/153] move(data/pi/*): Group `pi` files (#13826) Move `data.pi` to `data.pi.algebra` and `order.pilex` to `data.pi.lex`. --- src/algebra/group/pi.lean | 2 +- src/algebra/hom/equiv.lean | 2 +- src/algebra/ring/basic.lean | 2 +- src/data/fin/tuple/basic.lean | 3 ++- src/data/{pi.lean => pi/algebra.lean} | 0 src/{order/pilex.lean => data/pi/lex.lean} | 0 6 files changed, 5 insertions(+), 4 deletions(-) rename src/data/{pi.lean => pi/algebra.lean} (100%) rename src/{order/pilex.lean => data/pi/lex.lean} (100%) diff --git a/src/algebra/group/pi.lean b/src/algebra/group/pi.lean index b6fe4274ae7e5..6b7c5f6285253 100644 --- a/src/algebra/group/pi.lean +++ b/src/algebra/group/pi.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon, Patrick Massot -/ import algebra.hom.group_instances -import data.pi +import data.pi.algebra import data.set.function import data.set.pairwise import tactic.pi_instances diff --git a/src/algebra/hom/equiv.lean b/src/algebra/hom/equiv.lean index 856e1be3d76f9..37b8d917598e1 100644 --- a/src/algebra/hom/equiv.lean +++ b/src/algebra/hom/equiv.lean @@ -5,7 +5,7 @@ Authors: Johannes Hölzl, Callum Sutton, Yury Kudryashov -/ import algebra.group.type_tags import algebra.group_with_zero.basic -import data.pi +import data.pi.algebra /-! # Multiplicative and additive equivs diff --git a/src/algebra/ring/basic.lean b/src/algebra/ring/basic.lean index efefecdf77c51..f70453c26c8cc 100644 --- a/src/algebra/ring/basic.lean +++ b/src/algebra/ring/basic.lean @@ -6,7 +6,7 @@ Neil Strickland -/ import algebra.divisibility import algebra.regular.basic -import data.pi +import data.pi.algebra /-! # Properties and homomorphisms of semirings and rings diff --git a/src/data/fin/tuple/basic.lean b/src/data/fin/tuple/basic.lean index d392b685dfa32..c333485f05f3a 100644 --- a/src/data/fin/tuple/basic.lean +++ b/src/data/fin/tuple/basic.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn, Yury Kudryashov, Sébastien Gouëzel, Chris Hughes -/ import data.fin.basic -import order.pilex +import data.pi.lex + /-! # Operation on tuples diff --git a/src/data/pi.lean b/src/data/pi/algebra.lean similarity index 100% rename from src/data/pi.lean rename to src/data/pi/algebra.lean diff --git a/src/order/pilex.lean b/src/data/pi/lex.lean similarity index 100% rename from src/order/pilex.lean rename to src/data/pi/lex.lean From 475a533469a8f078f44213f9a97fefb5b9d9af11 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 3 May 2022 14:29:28 +0000 Subject: [PATCH 047/153] feat(topology/algebra/module/basic): A continuous linear functional is open (#13829) --- src/topology/algebra/module/basic.lean | 16 ++++++++++++++++ 1 file changed, 16 insertions(+) diff --git a/src/topology/algebra/module/basic.lean b/src/topology/algebra/module/basic.lean index 0be7182b4669a..8b654b46a1707 100644 --- a/src/topology/algebra/module/basic.lean +++ b/src/topology/algebra/module/basic.lean @@ -501,6 +501,9 @@ coe_injective.unique instance unique_of_right [subsingleton M₂] : unique (M₁ →SL[σ₁₂] M₂) := coe_injective.unique +lemma exists_ne_zero {f : M₁ →SL[σ₁₂] M₂} (hf : f ≠ 0) : ∃ x, f x ≠ 0 := +by { by_contra' h, exact hf (continuous_linear_map.ext h) } + section variables (R₁ M₁) @@ -1025,6 +1028,19 @@ end end ring +section division_monoid +variables {R M : Type*} + +/-- A nonzero continuous linear functional is open. -/ +protected lemma is_open_map_of_ne_zero [topological_space R] [division_ring R] + [has_continuous_sub R] [add_comm_group M] [topological_space M] [has_continuous_add M] + [module R M] [has_continuous_smul R M] (f : M →L[R] R) (hf : f ≠ 0) : is_open_map f := +let ⟨x, hx⟩ := exists_ne_zero hf in is_open_map.of_sections $ λ y, + ⟨λ a, y + (a - f y) • (f x)⁻¹ • x, continuous.continuous_at $ by continuity, + by simp, λ a, by simp [hx]⟩ + +end division_monoid + section smul_monoid -- The M's are used for semilinear maps, and the N's for plain linear maps From 0f8d7a99ca8bc9cafc10d4f853f76084e2215b45 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Tue, 3 May 2022 14:29:29 +0000 Subject: [PATCH 048/153] feat(order/omega_complete_partial_order): make `continuous_hom.prod.apply` continuous (#13833) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Previous it was defined as `apply : (α →𝒄 β) × α →o β` and the comment said that it would make sense to define it as a continuous function, but we need an instance for `α →𝒄 β` first. But then let’s just define that instance first, and then define `apply : (α →𝒄 β) × α →𝒄 β` as you would expect. Also rephrases `lemma ωSup_ωSup` differently now that `apply` is continuous. --- src/order/omega_complete_partial_order.lean | 61 +++++++++++++-------- 1 file changed, 38 insertions(+), 23 deletions(-) diff --git a/src/order/omega_complete_partial_order.lean b/src/order/omega_complete_partial_order.lean index 8a32d6fda2117..cbd3fb890ebb2 100644 --- a/src/order/omega_complete_partial_order.lean +++ b/src/order/omega_complete_partial_order.lean @@ -386,6 +386,13 @@ instance : omega_complete_partial_order (α × β) := le_ωSup := λ c i, ⟨le_ωSup (c.map order_hom.fst) i, le_ωSup (c.map order_hom.snd) i⟩ } +lemma ωSup_zip (c₀ : chain α) (c₁ : chain β) : + ωSup (c₀.zip c₁) = (ωSup c₀, ωSup c₁) := +begin + apply eq_of_forall_ge_iff, rintro ⟨z₁,z₂⟩, + simp [ωSup_le_iff, forall_and_distrib], +end + end prod open omega_complete_partial_order @@ -644,19 +651,6 @@ of_mono (order_hom.const _ x) (continuous_const x) instance [inhabited β] : inhabited (α →𝒄 β) := ⟨ const default ⟩ -namespace prod - -/-- The application of continuous functions as a monotone function. - -(It would make sense to make it a continuous function, but we are currently constructing a -`omega_complete_partial_order` instance for `α →𝒄 β`, and we cannot use it as the domain or image -of a continuous function before we do.) -/ -@[simps] -def apply : (α →𝒄 β) × α →o β := -{ to_fun := λ f, f.1 f.2, - monotone' := λ x y h, by dsimp; transitivity y.fst x.snd; [apply h.1, apply y.1.monotone h.2] } - -end prod /-- The map from continuous functions to monotone functions is itself a monotone function. -/ @[simps] @@ -704,18 +698,39 @@ instance : omega_complete_partial_order (α →𝒄 β) := omega_complete_partial_order.lift continuous_hom.to_mono continuous_hom.ωSup (λ x y h, h) (λ c, rfl) +namespace prod + +/-- The application of continuous functions as a continuous function. -/ +@[simps] +def apply : (α →𝒄 β) × α →𝒄 β := +{ to_fun := λ f, f.1 f.2, + monotone' := λ x y h, by {dsimp, transitivity y.fst x.snd; [apply h.1, apply y.1.monotone h.2]}, + cont := begin + intro c, + apply le_antisymm, + { apply ωSup_le, intros i, + dsimp, + rw (c _).fst.continuous, + apply ωSup_le, intros j, + apply le_ωSup_of_le (max i j), + apply apply_mono, + exact monotone_fst (order_hom.mono _ (le_max_left _ _)), + exact monotone_snd (order_hom.mono _ (le_max_right _ _)), }, + { apply ωSup_le, intros i, + apply le_ωSup_of_le i, + dsimp, + apply order_hom.mono _, + apply le_ωSup_of_le i, + reflexivity, } + end } + +end prod + lemma ωSup_def (c : chain (α →𝒄 β)) (x : α) : ωSup c x = continuous_hom.ωSup c x := rfl -lemma ωSup_ωSup (c₀ : chain (α →𝒄 β)) (c₁ : chain α) : - ωSup c₀ (ωSup c₁) = ωSup (continuous_hom.prod.apply.comp $ c₀.zip c₁) := -begin - apply eq_of_forall_ge_iff, intro z, - simp only [ωSup_le_iff, (c₀ _).continuous, chain.map_coe, to_mono_coe, coe_apply, - order_hom.omega_complete_partial_order_ωSup_coe, ωSup_def, forall_forall_merge, - chain.zip_coe, order_hom.prod_map_coe, order_hom.diag_coe, prod.map_mk, - order_hom.apply_coe, function.comp_app, prod.apply_coe, - order_hom.comp_coe, ωSup_apply, function.eval], -end +lemma ωSup_apply_ωSup (c₀ : chain (α →𝒄 β)) (c₁ : chain α) : + ωSup c₀ (ωSup c₁) = prod.apply (ωSup (c₀.zip c₁)) := +by simp [prod.apply_apply, prod.ωSup_zip] /-- A family of continuous functions yields a continuous family of functions. -/ @[simps] From 1c4d2b7ccfa5ce0d0921bb4f61aa27821db606b9 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 3 May 2022 14:29:30 +0000 Subject: [PATCH 049/153] feat(linear_algebra/matrix/trace): add `trace_conj_transpose` (#13888) --- src/linear_algebra/matrix/trace.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/src/linear_algebra/matrix/trace.lean b/src/linear_algebra/matrix/trace.lean index fa0e18ef89d24..f4953a34463cc 100644 --- a/src/linear_algebra/matrix/trace.lean +++ b/src/linear_algebra/matrix/trace.lean @@ -49,6 +49,10 @@ finset.smul_sum.symm @[simp] lemma trace_transpose (A : matrix n n R) : trace Aᵀ = trace A := rfl +@[simp] lemma trace_conj_transpose [star_add_monoid R] (A : matrix n n R) : + trace Aᴴ = star (trace A) := +(star_sum _ _).symm + variables (n α R) /-- `matrix.trace` as an `add_monoid_hom` -/ @[simps] From 40b595239b1441b80cb9b3eb3d20d7aa57c4e411 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 3 May 2022 14:29:31 +0000 Subject: [PATCH 050/153] doc(analysis/matrix): fix broken LaTeX (#13910) --- src/analysis/matrix.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/analysis/matrix.lean b/src/analysis/matrix.lean index cd12ceb35994d..750edd6c17c12 100644 --- a/src/analysis/matrix.lean +++ b/src/analysis/matrix.lean @@ -317,7 +317,7 @@ end linfty_op /-! ### The Frobenius norm -This is defined as $\|A\| = \sqrt{\sum_ij \|A_ij\|^2}$. +This is defined as $\|A\| = \sqrt{\sum_{i,j} \|A_{ij}\|^2}$. When the matrix is over the real or complex numbers, this norm is submultiplicative. -/ From 5c433d03f821fa9f111ea6ec6d9b5675e8d3a249 Mon Sep 17 00:00:00 2001 From: Stuart Presnell Date: Tue, 3 May 2022 16:28:59 +0000 Subject: [PATCH 051/153] feat(algebra/big_operators/basic): `prod_list_count` and `prod_list_count_of_subset` (#13370) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add `prod_list_count (l : list α) : l.prod = ∏ x in l.to_finset, (x ^ (l.count x))` and `prod_list_count_of_subset (l : list α) (s : finset α) (hs : l.to_finset ⊆ s) : l.prod = ∏ x in s, x ^ (l.count x)` as counterparts of `prod_multiset_count` and `prod_multiset_count_of_subset` (whose proofs are then golfed using the new lemmas). --- src/algebra/big_operators/basic.lean | 56 ++++++++++++++++++++-------- 1 file changed, 40 insertions(+), 16 deletions(-) diff --git a/src/algebra/big_operators/basic.lean b/src/algebra/big_operators/basic.lean index fda86da3426e1..4c941cb016542 100644 --- a/src/algebra/big_operators/basic.lean +++ b/src/algebra/big_operators/basic.lean @@ -933,41 +933,65 @@ lemma prod_range_one (f : ℕ → β) : ∏ k in range 1, f k = f 0 := by { rw [range_one], apply @prod_singleton β ℕ 0 f } -open multiset +open list -@[to_additive] lemma prod_multiset_map_count [decidable_eq α] (s : multiset α) +@[to_additive] lemma prod_list_map_count [decidable_eq α] (l : list α) {M : Type*} [comm_monoid M] (f : α → M) : - (s.map f).prod = ∏ m in s.to_finset, (f m) ^ (s.count m) := + (l.map f).prod = ∏ m in l.to_finset, (f m) ^ (l.count m) := begin - induction s using multiset.induction_on with a s ih, - { simp only [prod_const_one, count_zero, prod_zero, pow_zero, multiset.map_zero] }, - simp only [multiset.prod_cons, map_cons, to_finset_cons, ih], + induction l with a s IH, { simp only [map_nil, prod_nil, count_nil, pow_zero, prod_const_one] }, + simp only [list.map, list.prod_cons, to_finset_cons, IH], by_cases has : a ∈ s.to_finset, { rw [insert_eq_of_mem has, ← insert_erase has, prod_insert (not_mem_erase _ _), - prod_insert (not_mem_erase _ _), ← mul_assoc, count_cons_self, pow_succ], - congr' 1, refine prod_congr rfl (λ x hx, _), + prod_insert (not_mem_erase _ _), ← mul_assoc, count_cons_self, pow_succ], + congr' 1, + refine prod_congr rfl (λ x hx, _), rw [count_cons_of_ne (ne_of_mem_erase hx)] }, rw [prod_insert has, count_cons_self, count_eq_zero_of_not_mem (mt mem_to_finset.2 has), pow_one], - congr' 1, refine prod_congr rfl (λ x hx, _), + congr' 1, + refine prod_congr rfl (λ x hx, _), rw count_cons_of_ne, - rintro rfl, exact has hx + rintro rfl, + exact has hx, end +@[to_additive] +lemma prod_list_count [decidable_eq α] [comm_monoid α] (s : list α) : + s.prod = ∏ m in s.to_finset, m ^ (s.count m) := +by simpa using prod_list_map_count s id + +@[to_additive] +lemma prod_list_count_of_subset [decidable_eq α] [comm_monoid α] + (m : list α) (s : finset α) (hs : m.to_finset ⊆ s) : + m.prod = ∏ i in s, i ^ (m.count i) := +begin + rw prod_list_count, + refine prod_subset hs (λ x _ hx, _), + rw [mem_to_finset] at hx, + rw [count_eq_zero_of_not_mem hx, pow_zero], +end + +open multiset + +@[to_additive] lemma prod_multiset_map_count [decidable_eq α] (s : multiset α) + {M : Type*} [comm_monoid M] (f : α → M) : + (s.map f).prod = ∏ m in s.to_finset, (f m) ^ (s.count m) := +by { refine quot.induction_on s (λ l, _), simpa [prod_list_map_count l f] } + @[to_additive] lemma prod_multiset_count [decidable_eq α] [comm_monoid α] (s : multiset α) : s.prod = ∏ m in s.to_finset, m ^ (s.count m) := -by { convert prod_multiset_map_count s id, rw map_id } +by { convert prod_multiset_map_count s id, rw multiset.map_id } @[to_additive] lemma prod_multiset_count_of_subset [decidable_eq α] [comm_monoid α] (m : multiset α) (s : finset α) (hs : m.to_finset ⊆ s) : m.prod = ∏ i in s, i ^ (m.count i) := begin - rw prod_multiset_count, - apply prod_subset hs, - rintros x - hx, - rw [mem_to_finset] at hx, - rw [count_eq_zero_of_not_mem hx, pow_zero], + revert hs, + refine quot.induction_on m (λ l, _), + simp only [quot_mk_to_coe'', coe_prod, coe_count], + apply prod_list_count_of_subset l s end @[to_additive] lemma prod_mem_multiset [decidable_eq α] From 86887539c7b0ea357fd251800e81ff4a574170e1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 3 May 2022 16:29:00 +0000 Subject: [PATCH 052/153] feat(set_theory/game/basic): Inline instances (#13813) We also add a few missing instances. --- src/set_theory/game/basic.lean | 105 +++++++++++---------------------- 1 file changed, 33 insertions(+), 72 deletions(-) diff --git a/src/set_theory/game/basic.lean b/src/set_theory/game/basic.lean index 2626b6765f8c8..073c8fc00ec5c 100644 --- a/src/set_theory/game/basic.lean +++ b/src/set_theory/game/basic.lean @@ -48,12 +48,8 @@ open pgame namespace game -/-- The relation `x ≤ y` on games. -/ -def le : game → game → Prop := -quotient.lift₂ (λ x y, x ≤ y) (λ x₁ y₁ x₂ y₂ hx hy, propext (le_congr hx hy)) - instance : has_le game := -{ le := le } +⟨quotient.lift₂ (λ x y, x ≤ y) (λ x₁ y₁ x₂ y₂ hx hy, propext (le_congr hx hy))⟩ -- Adding `@[refl]` and `@[trans]` attributes here would override the ones on -- `preorder.le_refl` and `preorder.le_trans`, which breaks all non-`game` uses of `≤`! @@ -66,83 +62,46 @@ by { rintro ⟨x⟩ ⟨y⟩ ⟨z⟩, apply pgame.le_trans } theorem le_antisymm : ∀ x y : game, x ≤ y → y ≤ x → x = y := by { rintro ⟨x⟩ ⟨y⟩ h₁ h₂, apply quot.sound, exact ⟨h₁, h₂⟩ } -/-- The relation `x < y` on games. -/ --- We don't yet make this into an instance, because it will conflict with the (incorrect) notion --- of `<` provided by `partial_order` later. -def lt : game → game → Prop := -quotient.lift₂ (λ x y, x < y) (λ x₁ y₁ x₂ y₂ hx hy, propext (lt_congr hx hy)) +/-- This instance is incompatible with that provided by `game.partial_order`, which is why it's made +into a `def` instead. -/ +instance : has_lt game := +⟨quotient.lift₂ (λ x y, x < y) (λ x₁ y₁ x₂ y₂ hx hy, propext (lt_congr hx hy))⟩ -theorem not_le : ∀ {x y : game}, ¬ (x ≤ y) ↔ (lt y x) := +@[simp] theorem not_le : ∀ {x y : game}, ¬ x ≤ y ↔ y < x := by { rintro ⟨x⟩ ⟨y⟩, exact not_le } +@[simp] theorem not_lt : ∀ {x y : game}, ¬ x < y ↔ y ≤ x := +by { rintro ⟨x⟩ ⟨y⟩, exact not_lt } + instance : has_zero game := ⟨⟦0⟧⟩ instance : inhabited game := ⟨0⟩ instance : has_one game := ⟨⟦1⟧⟩ /-- The negation of `{L | R}` is `{-R | -L}`. -/ -def neg : game → game := -quot.lift (λ x, ⟦-x⟧) (λ x y h, quot.sound (@neg_congr x y h)) - instance : has_neg game := -{ neg := neg } +⟨quot.lift (λ x, ⟦-x⟧) (λ x y h, quot.sound (@neg_congr x y h))⟩ /-- The sum of `x = {xL | xR}` and `y = {yL | yR}` is `{xL + y, x + yL | xR + y, x + yR}`. -/ -def add : game → game → game := -quotient.lift₂ (λ x y : pgame, ⟦x + y⟧) (λ x₁ y₁ x₂ y₂ hx hy, quot.sound (pgame.add_congr hx hy)) - -instance : has_add game := ⟨add⟩ - -theorem add_assoc : ∀ (x y z : game), (x + y) + z = x + (y + z) := -begin - rintros ⟨x⟩ ⟨y⟩ ⟨z⟩, - apply quot.sound, - exact add_assoc_equiv -end +instance : has_add game := +⟨quotient.lift₂ (λ x y : pgame, ⟦x + y⟧) (λ x₁ y₁ x₂ y₂ hx hy, quot.sound (pgame.add_congr hx hy))⟩ instance : add_semigroup game.{u} := -{ add_assoc := add_assoc, +{ add_assoc := by { rintros ⟨x⟩ ⟨y⟩ ⟨z⟩, exact quot.sound add_assoc_equiv }, ..game.has_add } -theorem add_zero : ∀ (x : game), x + 0 = x := -begin - rintro ⟨x⟩, - apply quot.sound, - apply add_zero_equiv -end -theorem zero_add : ∀ (x : game), 0 + x = x := -begin - rintro ⟨x⟩, - apply quot.sound, - apply zero_add_equiv -end - instance : add_monoid game := -{ add_zero := add_zero, - zero_add := zero_add, +{ add_zero := by { rintro ⟨x⟩, exact quot.sound (add_zero_equiv x) }, + zero_add := by { rintro ⟨x⟩, exact quot.sound (zero_add_equiv x) }, ..game.has_zero, ..game.add_semigroup } -theorem add_left_neg : ∀ (x : game), (-x) + x = 0 := -begin - rintro ⟨x⟩, - apply quot.sound, - apply add_left_neg_equiv -end - instance : add_group game := -{ add_left_neg := add_left_neg, +{ add_left_neg := by { rintro ⟨x⟩, exact quot.sound (add_left_neg_equiv x) }, ..game.has_neg, ..game.add_monoid } -theorem add_comm : ∀ (x y : game), x + y = y + x := -begin - rintros ⟨x⟩ ⟨y⟩, - apply quot.sound, - exact add_comm_equiv -end - instance : add_comm_semigroup game := -{ add_comm := add_comm, +{ add_comm := by { rintros ⟨x⟩ ⟨y⟩, exact quot.sound add_comm_equiv }, ..game.add_semigroup } instance : add_comm_group game := @@ -150,10 +109,16 @@ instance : add_comm_group game := ..game.add_group } instance covariant_class_add_le : covariant_class game game (+) (≤) := -⟨begin rintro ⟨a⟩ ⟨b⟩ ⟨c⟩ h, exact @add_le_add_left _ _ _ _ b c h a end⟩ +⟨by { rintro ⟨a⟩ ⟨b⟩ ⟨c⟩ h, exact @add_le_add_left _ _ _ _ b c h a }⟩ instance covariant_class_swap_add_le : covariant_class game game (swap (+)) (≤) := -⟨begin rintro ⟨a⟩ ⟨b⟩ ⟨c⟩ h, exact @add_le_add_right _ _ _ _ b c h a end⟩ +⟨by { rintro ⟨a⟩ ⟨b⟩ ⟨c⟩ h, exact @add_le_add_right _ _ _ _ b c h a }⟩ + +instance covariant_class_add_lt : covariant_class game game (+) (<) := +⟨by { rintro ⟨a⟩ ⟨b⟩ ⟨c⟩ h, exact @add_lt_add_left _ _ _ _ b c h a }⟩ + +instance covariant_class_swap_add_lt : covariant_class game game (swap (+)) (<) := +⟨by { rintro ⟨a⟩ ⟨b⟩ ⟨c⟩ h, exact @add_lt_add_right _ _ _ _ b c h a }⟩ -- While it is very tempting to define a `partial_order` on games, and prove -- that games form an `ordered_add_comm_group`, it is a bit dangerous. @@ -168,17 +133,17 @@ instance covariant_class_swap_add_le : covariant_class game game (swap (+)) (≤ -- but do not actually mark them as instances, for safety. /-- The `<` operation provided by this partial order is not the usual `<` on games! -/ -def game_partial_order : partial_order game := +def partial_order : partial_order game := { le_refl := le_refl, le_trans := le_trans, le_antisymm := le_antisymm, ..game.has_le } /-- The `<` operation provided by this `ordered_add_comm_group` is not the usual `<` on games! -/ -def ordered_add_comm_group_game : ordered_add_comm_group game := +def ordered_add_comm_group : ordered_add_comm_group game := { add_le_add_left := @add_le_add_left _ _ _ game.covariant_class_add_le, ..game.add_comm_group, - ..game_partial_order } + ..game.partial_order } end game @@ -207,8 +172,8 @@ Hence we define them here. -/ /-- The product of `x = {xL | xR}` and `y = {yL | yR}` is `{xL*y + x*yL - xL*yL, xR*y + x*yR - xR*yR | xL*y + x*yR - xL*yR, x*yL + xR*y - xR*yL }`. -/ -def mul (x y : pgame) : pgame := -begin +instance : has_mul pgame.{u} := +⟨λ x y, begin induction x with xl xr xL xR IHxl IHxr generalizing y, induction y with yl yr yL yR IHyl IHyr, have y := mk yl yr yL yR, @@ -217,9 +182,7 @@ begin { exact IHxr i y + IHyr j - IHxr i (yR j) }, { exact IHxl i y + IHyr j - IHxl i (yR j) }, { exact IHxr i y + IHyl j - IHxr i (yL j) } -end - -instance : has_mul pgame := ⟨mul⟩ +end⟩ /-- An explicit description of the moves for Left in `x * y`. -/ def left_moves_mul (x y : pgame) : (x * y).left_moves @@ -557,11 +520,9 @@ def inv' : pgame → pgame inv_val L' R IHl' IHr, inv_val L' R IHl' IHr⟩ /-- The inverse of a surreal number in terms of the inverse on positive surreals. -/ -noncomputable def inv (x : pgame) : pgame := -by classical; exact -if x = 0 then 0 else if 0 < x then inv' x else inv' (-x) +noncomputable instance : has_inv pgame := +⟨by { classical, exact λ x, if x = 0 then 0 else if 0 < x then inv' x else inv' (-x) }⟩ -noncomputable instance : has_inv pgame := ⟨inv⟩ noncomputable instance : has_div pgame := ⟨λ x y, x * y⁻¹⟩ end pgame From 7d28753753500766d0b0b2a36a3b3ed37e18cb57 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 3 May 2022 16:29:01 +0000 Subject: [PATCH 053/153] chore(normed_space/weak_dual): generalize `normed_group` to `semi_normed_group` (#13914) This almost halves the time this file takes to build, and is more general too. --- src/analysis/normed_space/weak_dual.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/analysis/normed_space/weak_dual.lean b/src/analysis/normed_space/weak_dual.lean index 2cdca8a61ce5f..4bac47c508a80 100644 --- a/src/analysis/normed_space/weak_dual.lean +++ b/src/analysis/normed_space/weak_dual.lean @@ -83,7 +83,7 @@ by the dual-norm (i.e. the operator-norm). open normed_space variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] -variables {E : Type*} [normed_group E] [normed_space 𝕜 E] +variables {E : Type*} [semi_normed_group E] [normed_space 𝕜 E] /-- For normed spaces `E`, there is a canonical map `dual 𝕜 E → weak_dual 𝕜 E` (the "identity" mapping). It is a linear equivalence. -/ From bd1d93596306387096bd9ea14aac11f6514c3e95 Mon Sep 17 00:00:00 2001 From: MichaelStollBayreuth Date: Tue, 3 May 2022 18:19:53 +0000 Subject: [PATCH 054/153] feat(number_theory/legendre_symbol/): add some lemmas (#13831) This adds essentially two lemmas on quadratic characters: * `quadratic_char_neg_one_iff_not_is_square`, which says that the quadratic character takes the value `-1` exactly on non-squares, and * `quadratic_char_number_of_sqrts`. which says that the number of square roots of `a : F` is `quadratic_char F a + 1`. It also adds the corresponding statements, `legendre_sym_eq_neg_one_iff` and `legendre_sym_number_of_sqrts`, for the Legendre symbol. --- .../legendre_symbol/quadratic_char.lean | 75 +++++++++++++++++-- .../quadratic_reciprocity.lean | 18 +++++ 2 files changed, 88 insertions(+), 5 deletions(-) diff --git a/src/number_theory/legendre_symbol/quadratic_char.lean b/src/number_theory/legendre_symbol/quadratic_char.lean index 98ecd288f77c1..6ccd8acff84f8 100644 --- a/src/number_theory/legendre_symbol/quadratic_char.lean +++ b/src/number_theory/legendre_symbol/quadratic_char.lean @@ -58,6 +58,16 @@ begin exact char_p.neg_one_ne_one _ (ring_char F), end +/-- Characteristic `≠ 2` implies that `-a ≠ a` when `a ≠ 0`. -/ +lemma neg_ne_self_of_char_ne_two (hF : ring_char F ≠ 2) {a : F} (ha : a ≠ 0) : a ≠ -a := +begin + intro hf, + rw [eq_neg_iff_add_eq_zero, (by ring : a + a = 2 * a), mul_eq_zero] at hf, + have h := mt eq_neg_iff_add_eq_zero.mpr (neg_one_ne_one_of_char_ne_two hF).symm, + norm_num at h, + exact or.dcases_on hf (λ (hf : 2 = 0), h hf) (λ (hf : a = 0), ha hf), +end + /-- If `F` has odd characteristic, then for nonzero `a : F`, we have that `a ^ (#F / 2) = ±1`. -/ lemma pow_dichotomy (hF : ring_char F ≠ 2) {a : F} (ha : a ≠ 0) : a^(fintype.card F / 2) = 1 ∨ a^(fintype.card F / 2) = -1 := @@ -263,14 +273,69 @@ lemma quadratic_char_dichotomy {a : F} (ha : a ≠ 0) : quadratic_char F a = 1 ∨ quadratic_char F a = -1 := (sq_eq_one_iff (quadratic_char F a)).mp (quadratic_char_sq_one ha) +/-- A variant -/ +lemma quadratic_char_eq_neg_one_iff_not_one {a : F} (ha : a ≠ 0) : + quadratic_char F a = -1 ↔ ¬ quadratic_char F a = 1 := +begin + split, + { intro h, + rw h, + norm_num, }, + { exact λ h₂, (or_iff_right h₂).mp (quadratic_char_dichotomy ha), } +end + +/-- For `a : F`, `quadratic_char F a = -1 ↔ ¬ is_square a`. -/ +lemma quadratic_char_neg_one_iff_not_is_square {a : F} : + quadratic_char F a = -1 ↔ ¬ is_square a := +begin + by_cases ha : a = 0, + { simp only [ha, is_square_zero, quadratic_char_zero, zero_eq_neg, one_ne_zero, not_true], }, + { rw quadratic_char_eq_neg_one_iff_not_one ha, + exact not_iff_not.mpr (quadratic_char_one_iff_is_square ha), }, +end + /-- If `F` has odd characteristic, then `quadratic_char F` takes the value `-1`. -/ lemma quadratic_char_exists_neg_one (hF : ring_char F ≠ 2) : ∃ a, quadratic_char F a = -1 := +Exists.dcases_on (finite_field.exists_nonsquare hF) + (λ (b : F) (h₁ : ¬is_square b), ⟨b, (quadratic_char_neg_one_iff_not_is_square.mpr h₁)⟩) + +/-- The number of solutions to `x^2 = a` is determined by the quadratic character. -/ +lemma quadratic_char_card_sqrts (hF : ring_char F ≠ 2) (a : F) : + ↑{x : F | x^2 = a}.to_finset.card = quadratic_char F a + 1 := begin - cases (finite_field.exists_nonsquare hF) with b h₁, - have hb : b ≠ 0 := by { intro hf, rw hf at h₁, exact h₁ (is_square_zero F), }, - use b, - simp only [quadratic_char, hb, if_false, ite_eq_right_iff], - tauto, + -- we consider the cases `a = 0`, `a` is a nonzero square and `a` is a nonsquare in turn + by_cases h₀ : a = 0, + { simp only [h₀, pow_eq_zero_iff, nat.succ_pos', int.coe_nat_succ, int.coe_nat_zero, zero_add, + quadratic_char_zero, add_zero, set.set_of_eq_eq_singleton, set.to_finset_card, + set.card_singleton], }, + { set s := {x : F | x^2 = a}.to_finset with hs, + by_cases h : is_square a, + { rw (quadratic_char_one_iff_is_square h₀).mpr h, + rcases h with ⟨b, h⟩, + have hb₀ : b ≠ 0 := by { intro hb, rw [hb, mul_zero] at h, exact h₀ h, }, + have h₁ : s = [b, -b].to_finset := by + { ext x, + simp only [finset.mem_filter, finset.mem_univ, true_and, list.to_finset_cons, + list.to_finset_nil, insert_emptyc_eq, finset.mem_insert, finset.mem_singleton], + rw ← pow_two at h, + rw hs, + simp only [set.mem_to_finset, set.mem_set_of_eq], + rw h, + split, + { exact eq_or_eq_neg_of_sq_eq_sq _ _, }, + { rintro (h₂ | h₂); rw h₂, + simp only [neg_sq], }, }, + simp only [h₁, finset.card_doubleton (finite_field.neg_ne_self_of_char_ne_two hF hb₀), + list.to_finset_cons, list.to_finset_nil, insert_emptyc_eq, int.coe_nat_succ, + int.coe_nat_zero, zero_add], }, + { rw quadratic_char_neg_one_iff_not_is_square.mpr h, + simp only [int.coe_nat_eq_zero, finset.card_eq_zero, set.to_finset_card, + fintype.card_of_finset, set.mem_set_of_eq, add_left_neg], + ext x, + simp only [iff_false, finset.mem_filter, finset.mem_univ, true_and, finset.not_mem_empty], + rw is_square_iff_exists_sq at h, + push_neg at h, + exact (h x).symm, }, }, end open_locale big_operators diff --git a/src/number_theory/legendre_symbol/quadratic_reciprocity.lean b/src/number_theory/legendre_symbol/quadratic_reciprocity.lean index 0cefbb4b86f04..3874d64bdbc83 100644 --- a/src/number_theory/legendre_symbol/quadratic_reciprocity.lean +++ b/src/number_theory/legendre_symbol/quadratic_reciprocity.lean @@ -62,6 +62,7 @@ begin refine ⟨units.mk0 y hy, _⟩, simp, } end +-- The following is used by number_theory/zsqrtd/gaussian_int.lean and archive/imo/imo2008_q3.lean lemma exists_sq_eq_neg_one_iff : (∃ y : zmod p, y ^ 2 = -1) ↔ p % 4 ≠ 3 := begin @@ -148,6 +149,10 @@ lemma legendre_sym_eq_one_or_neg_one (p : ℕ) [fact p.prime] (a : ℤ) (ha : (a legendre_sym p a = 1 ∨ legendre_sym p a = -1 := quadratic_char_dichotomy ha +lemma legendre_sym_eq_neg_one_iff_not_one {a : ℤ} (ha : (a : zmod p) ≠ 0) : + legendre_sym p a = -1 ↔ ¬ legendre_sym p a = 1 := +quadratic_char_eq_neg_one_iff_not_one ha + /-- The Legendre symbol of `p` and `a` is zero iff `p ∣ a`. -/ lemma legendre_sym_eq_zero_iff (p : ℕ) [fact p.prime] (a : ℤ) : legendre_sym p a = 0 ↔ (a : zmod p) = 0 := @@ -222,6 +227,19 @@ lemma legendre_sym_eq_one_iff {a : ℤ} (ha0 : (a : zmod p) ≠ 0) : legendre_sym p a = 1 ↔ is_square (a : zmod p) := quadratic_char_one_iff_is_square ha0 +/-- `legendre_sym p a = -1` iff`a` is a nonsquare mod `p`. -/ +lemma legendre_sym_eq_neg_one_iff {a : ℤ} : + legendre_sym p a = -1 ↔ ¬ is_square (a : zmod p) := +quadratic_char_neg_one_iff_not_is_square + +/-- The number of square roots of `a` modulo `p` is determined by the Legendre symbol. -/ +lemma legendre_sym_card_sqrts (hp : p ≠ 2) (a : ℤ) : + ↑{x : zmod p | x^2 = a}.to_finset.card = legendre_sym p a + 1 := +begin + have h : ring_char (zmod p) ≠ 2 := by { rw ring_char_zmod_n, exact hp, }, + exact quadratic_char_card_sqrts h a, +end + open_locale big_operators lemma eisenstein_lemma (hp : p ≠ 2) {a : ℕ} (ha1 : a % 2 = 1) (ha0 : (a : zmod p) ≠ 0) : From 16157f26bde58e04ae3babf40ef89c8ad147dd1c Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 3 May 2022 18:19:54 +0000 Subject: [PATCH 055/153] chore(topology/continuous_function/bounded): generalize from `normed_*` to `semi_normed_*` (#13915) Every single lemma in this file generalized, apart from the ones that transferred a `normed_*` instance which obviously need the stronger assumption. `dist_zero_of_empty` was the only lemma that actually needed reproving from scratch, all the other affected proofs are just split between two instances. --- src/topology/continuous_function/bounded.lean | 93 ++++++++++++------- 1 file changed, 58 insertions(+), 35 deletions(-) diff --git a/src/topology/continuous_function/bounded.lean b/src/topology/continuous_function/bounded.lean index 3c92c051eed18..7ecf9de616238 100644 --- a/src/topology/continuous_function/bounded.lean +++ b/src/topology/continuous_function/bounded.lean @@ -32,8 +32,8 @@ When possible, instead of parametrizing results over `(f : α →ᵇ β)`, you should parametrize over `(F : Type*) [bounded_continuous_map_class F α β] (f : F)`. When you extend this structure, make sure to extend `bounded_continuous_map_class`. -/ -structure bounded_continuous_function - (α : Type u) (β : Type v) [topological_space α] [metric_space β] extends continuous_map α β : +structure bounded_continuous_function (α : Type u) (β : Type v) + [topological_space α] [pseudo_metric_space β] extends continuous_map α β : Type (max u v) := (map_bounded' : ∃ C, ∀ x y, dist (to_fun x) (to_fun y) ≤ C) @@ -42,7 +42,7 @@ localized "infixr ` →ᵇ `:25 := bounded_continuous_function" in bounded_conti /-- `bounded_continuous_map_class F α β` states that `F` is a type of bounded continuous maps. You should also extend this typeclass when you extend `bounded_continuous_function`. -/ -class bounded_continuous_map_class (F α β : Type*) [topological_space α] [metric_space β] +class bounded_continuous_map_class (F α β : Type*) [topological_space α] [pseudo_metric_space β] extends continuous_map_class F α β := (map_bounded (f : F) : ∃ C, ∀ x y, dist (f x) (f y) ≤ C) @@ -50,7 +50,7 @@ export bounded_continuous_map_class (map_bounded) namespace bounded_continuous_function section basics -variables [topological_space α] [metric_space β] [metric_space γ] +variables [topological_space α] [pseudo_metric_space β] [pseudo_metric_space γ] variables {f g : α →ᵇ β} {x : α} {C : ℝ} instance : bounded_continuous_map_class (α →ᵇ β) α β := @@ -170,16 +170,19 @@ lemma dist_lt_iff_of_nonempty_compact [nonempty α] [compact_space α] : dist f g < C ↔ ∀x:α, dist (f x) (g x) < C := ⟨λ w x, lt_of_le_of_lt (dist_coe_le_dist x) w, dist_lt_of_nonempty_compact⟩ -/-- The type of bounded continuous functions, with the uniform distance, is a metric space. -/ -instance : metric_space (α →ᵇ β) := +/-- The type of bounded continuous functions, with the uniform distance, is a pseudometric space. -/ +instance : pseudo_metric_space (α →ᵇ β) := { dist_self := λ f, le_antisymm ((dist_le le_rfl).2 $ λ x, by simp) dist_nonneg', - eq_of_dist_eq_zero := λ f g hfg, by ext x; exact - eq_of_dist_eq_zero (le_antisymm (hfg ▸ dist_coe_le_dist _) dist_nonneg), dist_comm := λ f g, by simp [dist_eq, dist_comm], dist_triangle := λ f g h, (dist_le (add_nonneg dist_nonneg' dist_nonneg')).2 $ λ x, le_trans (dist_triangle _ _ _) (add_le_add (dist_coe_le_dist _) (dist_coe_le_dist _)) } +/-- The type of bounded continuous functions, with the uniform distance, is a metric space. -/ +instance {α β} [topological_space α] [metric_space β] : metric_space (α →ᵇ β) := +{ eq_of_dist_eq_zero := λ f g hfg, by ext x; exact + eq_of_dist_eq_zero (le_antisymm (hfg ▸ dist_coe_le_dist _) dist_nonneg) } + lemma nndist_eq : nndist f g = Inf {C | ∀ x : α, nndist (f x) (g x) ≤ C} := subtype.ext $ dist_eq.trans $ begin rw [nnreal.coe_Inf, nnreal.coe_image], @@ -194,7 +197,7 @@ dist_coe_le_dist x /-- On an empty space, bounded continuous functions are at distance 0 -/ lemma dist_zero_of_empty [is_empty α] : dist f g = 0 := -dist_eq_zero.2 (eq_of_empty f g) +by rw [(ext is_empty_elim : f = g), dist_self] lemma dist_eq_supr : dist f g = ⨆ x : α, dist (f x) (g x) := begin @@ -396,7 +399,7 @@ end extend end basics section arzela_ascoli -variables [topological_space α] [compact_space α] [metric_space β] +variables [topological_space α] [compact_space α] [pseudo_metric_space β] variables {f g : α →ᵇ β} {x : α} {C : ℝ} /- Arzela-Ascoli theorem asserts that, on a compact space, a set of functions sharing @@ -494,7 +497,7 @@ end /-- Third (main) version, with pointwise equicontinuity and range in a compact subset, but without closedness. The closure is then compact -/ -theorem arzela_ascoli +theorem arzela_ascoli [t2_space β] (s : set β) (hs : is_compact s) (A : set (α →ᵇ β)) (in_s : ∀(f : α →ᵇ β) (x : α), f ∈ A → f x ∈ s) @@ -524,7 +527,7 @@ arzela_ascoli₂ s hs (closure A) is_closed_closure instance is when the source space is a metric space, and there is a fixed modulus of continuity for all the functions in the set A -/ -lemma equicontinuous_of_continuity_modulus {α : Type u} [metric_space α] +lemma equicontinuous_of_continuity_modulus {α : Type u} [pseudo_metric_space α] (b : ℝ → ℝ) (b_lim : tendsto b (𝓝 0) (𝓝 0)) (A : set (α →ᵇ β)) (H : ∀(x y:α) (f : α →ᵇ β), f ∈ A → dist (f x) (f y) ≤ b (dist x y)) @@ -548,7 +551,7 @@ end arzela_ascoli section has_one -variables [topological_space α] [metric_space β] [has_one β] +variables [topological_space α] [pseudo_metric_space β] [has_one β] @[to_additive] instance : has_one (α →ᵇ β) := ⟨const α 1⟩ @@ -577,7 +580,7 @@ names (for example, `coe_mul`) to conflict with later lemma names for normed rin trivial inconvenience, but in any case there are no obvious applications of the multiplicative version. -/ -variables [topological_space α] [metric_space β] [add_monoid β] +variables [topological_space α] [pseudo_metric_space β] [add_monoid β] variables [has_lipschitz_add β] variables (f g : α →ᵇ β) {x : α} {C : ℝ} @@ -649,7 +652,7 @@ end has_lipschitz_add section comm_has_lipschitz_add -variables [topological_space α] [metric_space β] [add_comm_monoid β] [has_lipschitz_add β] +variables [topological_space α] [pseudo_metric_space β] [add_comm_monoid β] [has_lipschitz_add β] @[to_additive] instance : add_comm_monoid (α →ᵇ β) := { add_comm := assume f g, by ext; simp [add_comm], @@ -672,7 +675,7 @@ section normed_group continuous functions from α to β inherits a normed group structure, by using pointwise operations and checking that they are compatible with the uniform distance. -/ -variables [topological_space α] [normed_group β] +variables [topological_space α] [semi_normed_group β] variables (f g : α →ᵇ β) {x : α} {C : ℝ} instance : has_norm (α →ᵇ β) := ⟨λu, dist u 0⟩ @@ -753,12 +756,12 @@ le_antisymm (norm_const_le b) $ h.elim $ λ x, (const α b).norm_coe_le_norm x /-- Constructing a bounded continuous function from a uniformly bounded continuous function taking values in a normed group. -/ -def of_normed_group {α : Type u} {β : Type v} [topological_space α] [normed_group β] +def of_normed_group {α : Type u} {β : Type v} [topological_space α] [semi_normed_group β] (f : α → β) (Hf : continuous f) (C : ℝ) (H : ∀x, ∥f x∥ ≤ C) : α →ᵇ β := ⟨⟨λn, f n, Hf⟩, ⟨_, dist_le_two_norm' H⟩⟩ @[simp] lemma coe_of_normed_group - {α : Type u} {β : Type v} [topological_space α] [normed_group β] + {α : Type u} {β : Type v} [topological_space α] [semi_normed_group β] (f : α → β) (Hf : continuous f) (C : ℝ) (H : ∀x, ∥f x∥ ≤ C) : (of_normed_group f Hf C H : α → β) = f := rfl @@ -769,16 +772,16 @@ lemma norm_of_normed_group_le {f : α → β} (hfc : continuous f) {C : ℝ} (hC /-- Constructing a bounded continuous function from a uniformly bounded function on a discrete space, taking values in a normed group -/ def of_normed_group_discrete {α : Type u} {β : Type v} - [topological_space α] [discrete_topology α] [normed_group β] + [topological_space α] [discrete_topology α] [semi_normed_group β] (f : α → β) (C : ℝ) (H : ∀x, norm (f x) ≤ C) : α →ᵇ β := of_normed_group f continuous_of_discrete_topology C H @[simp] lemma coe_of_normed_group_discrete - {α : Type u} {β : Type v} [topological_space α] [discrete_topology α] [normed_group β] + {α : Type u} {β : Type v} [topological_space α] [discrete_topology α] [semi_normed_group β] (f : α → β) (C : ℝ) (H : ∀x, ∥f x∥ ≤ C) : (of_normed_group_discrete f C H : α → β) = f := rfl -/-- Taking the pointwise norm of a bounded continuous function with values in a `normed_group`, +/-- Taking the pointwise norm of a bounded continuous function with values in a `semi_normed_group`, yields a bounded continuous function with values in ℝ. -/ def norm_comp : α →ᵇ ℝ := f.comp norm lipschitz_with_one_norm @@ -834,9 +837,12 @@ instance : add_comm_group (α →ᵇ β) := fun_like.coe_injective.add_comm_group _ coe_zero coe_add coe_neg coe_sub (λ _ _, coe_nsmul _ _) (λ _ _, coe_zsmul _ _) -instance : normed_group (α →ᵇ β) := +instance : semi_normed_group (α →ᵇ β) := { dist_eq := λ f g, by simp only [norm_eq, dist_eq, dist_eq_norm, sub_apply] } +instance {α β} [topological_space α] [normed_group β] : normed_group (α →ᵇ β) := +{ ..bounded_continuous_function.semi_normed_group } + lemma nnnorm_def : ∥f∥₊ = nndist f 0 := rfl lemma nnnorm_coe_le_nnnorm (x : α) : ∥f x∥₊ ≤ ∥f∥₊ := norm_coe_le_norm _ _ @@ -879,7 +885,7 @@ functions from `α` to `β` inherits a so-called `has_bounded_smul` structure (i `has_continuous_mul` structure, which is the mathlib formulation of being a topological module), by using pointwise operations and checking that they are compatible with the uniform distance. -/ -variables {𝕜 : Type*} [pseudo_metric_space 𝕜] [topological_space α] [metric_space β] +variables {𝕜 : Type*} [pseudo_metric_space 𝕜] [topological_space α] [pseudo_metric_space β] section has_scalar variables [has_zero 𝕜] [has_zero β] [has_scalar 𝕜 β] [has_bounded_smul 𝕜 β] @@ -973,7 +979,7 @@ continuous functions from `α` to `β` inherits a normed space structure, by usi pointwise operations and checking that they are compatible with the uniform distance. -/ variables {𝕜 : Type*} -variables [topological_space α] [normed_group β] +variables [topological_space α] [semi_normed_group β] variables {f g : α →ᵇ β} {x : α} {C : ℝ} instance [normed_field 𝕜] [normed_space 𝕜 β] : normed_space 𝕜 (α →ᵇ β) := ⟨λ c f, begin @@ -982,7 +988,7 @@ instance [normed_field 𝕜] [normed_space 𝕜 β] : normed_space 𝕜 (α → (mul_le_mul_of_nonneg_left (f.norm_coe_le_norm _) (norm_nonneg _))) end⟩ variables [nondiscrete_normed_field 𝕜] [normed_space 𝕜 β] -variables [normed_group γ] [normed_space 𝕜 γ] +variables [semi_normed_group γ] [normed_space 𝕜 γ] variables (α) -- TODO does this work in the `has_bounded_smul` setting, too? @@ -1023,11 +1029,12 @@ variables [topological_space α] {R : Type*} section non_unital -variables [non_unital_normed_ring R] +section semi_normed +variables [non_unital_semi_normed_ring R] instance : has_mul (α →ᵇ R) := { mul := λ f g, of_normed_group (f * g) (f.continuous.mul g.continuous) (∥f∥ * ∥g∥) $ λ x, - le_trans (non_unital_normed_ring.norm_mul (f x) (g x)) $ + le_trans (norm_mul_le (f x) (g x)) $ mul_le_mul (f.norm_coe_le_norm x) (g.norm_coe_le_norm x) (norm_nonneg _) (norm_nonneg _) } @[simp] lemma coe_mul (f g : α →ᵇ R) : ⇑(f * g) = f * g := rfl @@ -1037,13 +1044,21 @@ instance : non_unital_ring (α →ᵇ R) := fun_like.coe_injective.non_unital_ring _ coe_zero coe_add coe_mul coe_neg coe_sub (λ _ _, coe_nsmul _ _) (λ _ _, coe_zsmul _ _) -instance : non_unital_normed_ring (α →ᵇ R) := +instance : non_unital_semi_normed_ring (α →ᵇ R) := { norm_mul := λ f g, norm_of_normed_group_le _ (mul_nonneg (norm_nonneg _) (norm_nonneg _)) _, + .. bounded_continuous_function.semi_normed_group } + +end semi_normed + +instance [non_unital_normed_ring R] : non_unital_normed_ring (α →ᵇ R) := +{ .. bounded_continuous_function.non_unital_semi_normed_ring, .. bounded_continuous_function.normed_group } end non_unital -variables [normed_ring R] +section semi_normed + +variables [semi_normed_ring R] @[simp] lemma coe_npow_rec (f : α →ᵇ R) : ∀ n, ⇑(npow_rec n f) = f ^ n | 0 := by rw [npow_rec, pow_zero, coe_one] @@ -1063,7 +1078,12 @@ fun_like.coe_injective.ring _ coe_zero coe_one coe_add coe_mul coe_neg coe_sub (λ _ _, coe_zsmul _ _) (λ _ _, coe_pow _ _) -instance : normed_ring (α →ᵇ R) := +instance : semi_normed_ring (α →ᵇ R) := +{ ..bounded_continuous_function.non_unital_semi_normed_ring } + +end semi_normed + +instance [normed_ring R] : normed_ring (α →ᵇ R) := { ..bounded_continuous_function.non_unital_normed_ring } end normed_ring @@ -1076,13 +1096,16 @@ In this section, if `R` is a normed commutative ring, then we show that the spac continuous functions from `α` to `R` inherits a normed commutative ring structure, by using pointwise operations and checking that they are compatible with the uniform distance. -/ -variables [topological_space α] {R : Type*} [normed_comm_ring R] +variables [topological_space α] {R : Type*} -instance : comm_ring (α →ᵇ R) := +instance [semi_normed_comm_ring R] : comm_ring (α →ᵇ R) := { mul_comm := λ f₁ f₂, ext $ λ x, mul_comm _ _, .. bounded_continuous_function.ring } -instance : normed_comm_ring (α →ᵇ R) := +instance [semi_normed_comm_ring R] : semi_normed_comm_ring (α →ᵇ R) := +{ .. bounded_continuous_function.comm_ring, .. bounded_continuous_function.semi_normed_group } + +instance [normed_comm_ring R] : normed_comm_ring (α →ᵇ R) := { .. bounded_continuous_function.comm_ring, .. bounded_continuous_function.normed_group } end normed_comm_ring @@ -1096,7 +1119,7 @@ continuous functions from `α` to `γ` inherits a normed algebra structure, by u pointwise operations and checking that they are compatible with the uniform distance. -/ variables {𝕜 : Type*} [normed_field 𝕜] -variables [topological_space α] [normed_group β] [normed_space 𝕜 β] +variables [topological_space α] [semi_normed_group β] [normed_space 𝕜 β] variables [normed_ring γ] [normed_algebra 𝕜 γ] variables {f g : α →ᵇ γ} {x : α} {c : 𝕜} @@ -1182,7 +1205,7 @@ completeness is guaranteed when `β` is complete (see section normed_group variables {𝕜 : Type*} [normed_field 𝕜] [star_ring 𝕜] -variables [topological_space α] [normed_group β] [star_add_monoid β] [normed_star_group β] +variables [topological_space α] [semi_normed_group β] [star_add_monoid β] [normed_star_group β] variables [normed_space 𝕜 β] [star_module 𝕜 β] instance : star_add_monoid (α →ᵇ β) := From 5cfb8dbcc9965bd6270947f4643172c20979245a Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Tue, 3 May 2022 19:18:51 +0000 Subject: [PATCH 056/153] refactor(ring_theory/jacobson_ideal): generalize lemmas to non-commutative rings (#13865) The main change here is that the order of multiplication has been adjusted slightly in `mem_jacobson_iff`and `exists_mul_sub_mem_of_sub_one_mem_jacobson`. In the commutative case this doesn't matter anyway. All the other changes are just moving lemmas between sections, the statements of no lemmas other than those two have been changed. No lemmas have been added or removed. The lemmas about `is_unit` and quotients don't generalize as easily, so I've not attempted to touch those; that would require some mathematical insight, which is out of scope for this PR! --- src/ring_theory/jacobson_ideal.lean | 99 ++++++++++++++++------------- src/ring_theory/nakayama.lean | 4 +- 2 files changed, 57 insertions(+), 46 deletions(-) diff --git a/src/ring_theory/jacobson_ideal.lean b/src/ring_theory/jacobson_ideal.lean index e8eb9f9819f39..a210e27453396 100644 --- a/src/ring_theory/jacobson_ideal.lean +++ b/src/ring_theory/jacobson_ideal.lean @@ -41,13 +41,15 @@ Jacobson, Jacobson radical, Local Ideal universes u v namespace ideal -variables {R : Type u} [comm_ring R] {I : ideal R} -variables {S : Type v} [comm_ring S] +variables {R : Type u} {S : Type v} open_locale polynomial section jacobson -/-- The Jacobson radical of `I` is the infimum of all maximal ideals containing `I`. -/ +section ring +variables [ring R] [ring S] {I : ideal R} + +/-- The Jacobson radical of `I` is the infimum of all maximal (left) ideals containing `I`. -/ def jacobson (I : ideal R) : ideal R := Inf {J : ideal R | I ≤ J ∧ is_maximal J} @@ -57,12 +59,6 @@ lemma le_jacobson : I ≤ jacobson I := @[simp] lemma jacobson_idem : jacobson (jacobson I) = jacobson I := le_antisymm (Inf_le_Inf (λ J hJ, ⟨Inf_le hJ, hJ.2⟩)) le_jacobson -lemma radical_le_jacobson : radical I ≤ jacobson I := -le_Inf (λ J hJ, (radical_eq_Inf I).symm ▸ Inf_le ⟨hJ.left, is_maximal.is_prime hJ.right⟩) - -lemma eq_radical_of_eq_jacobson : jacobson I = I → radical I = I := -λ h, le_antisymm (le_trans radical_le_jacobson (le_of_eq h)) le_radical - @[simp] lemma jacobson_top : jacobson (⊤ : ideal R) = ⊤ := eq_top_iff.2 le_jacobson @@ -84,42 +80,31 @@ instance jacobson.is_maximal [H : is_maximal I] : is_maximal (jacobson I) := ⟨⟨λ htop, H.1.1 (jacobson_eq_top_iff.1 htop), λ J hJ, H.1.2 _ (lt_of_le_of_lt le_jacobson hJ)⟩⟩ -theorem mem_jacobson_iff {x : R} : x ∈ jacobson I ↔ ∀ y, ∃ z, x * y * z + z - 1 ∈ I := +theorem mem_jacobson_iff {x : R} : x ∈ jacobson I ↔ ∀ y, ∃ z, z * y * x + z - 1 ∈ I := ⟨λ hx y, classical.by_cases - (assume hxy : I ⊔ span {x * y + 1} = ⊤, + (assume hxy : I ⊔ span {y * x + 1} = ⊤, let ⟨p, hpi, q, hq, hpq⟩ := submodule.mem_sup.1 ((eq_top_iff_one _).1 hxy) in - let ⟨r, hr⟩ := mem_span_singleton.1 hq in - ⟨r, by rw [← one_mul r, ← mul_assoc, ← add_mul, mul_one, ← hr, ← hpq, ← neg_sub, - add_sub_cancel]; exact I.neg_mem hpi⟩) - (assume hxy : I ⊔ span {x * y + 1} ≠ ⊤, + let ⟨r, hr⟩ := mem_span_singleton'.1 hq in + ⟨r, by rw [mul_assoc, ←mul_add_one, hr, ← hpq, ← neg_sub, add_sub_cancel]; exact I.neg_mem hpi⟩) + (assume hxy : I ⊔ span {y * x + 1} ≠ ⊤, let ⟨M, hm1, hm2⟩ := exists_le_maximal _ hxy in suffices x ∉ M, from (this $ mem_Inf.1 hx ⟨le_trans le_sup_left hm2, hm1⟩).elim, - λ hxm, hm1.1.1 $ (eq_top_iff_one _).2 $ add_sub_cancel' (x * y) 1 ▸ M.sub_mem - (le_sup_right.trans hm2 $ mem_span_singleton.2 dvd_rfl) - (M.mul_mem_right _ hxm)), + λ hxm, hm1.1.1 $ (eq_top_iff_one _).2 $ add_sub_cancel' (y * x) 1 ▸ M.sub_mem + (le_sup_right.trans hm2 $ subset_span rfl) + (M.mul_mem_left _ hxm)), λ hx, mem_Inf.2 $ λ M ⟨him, hm⟩, classical.by_contradiction $ λ hxm, - let ⟨y, hy⟩ := hm.exists_inv hxm, ⟨z, hz⟩ := hx (-y) in - hm.1.1 $ (eq_top_iff_one _).2 $ sub_sub_cancel (x * -y * z + z) 1 ▸ M.sub_mem - (by { rw [← one_mul z, ← mul_assoc, ← add_mul, mul_one, mul_neg, neg_add_eq_sub, - ← neg_sub, neg_mul, neg_mul_eq_mul_neg, mul_comm x y, mul_comm _ (- z)], - rcases hy with ⟨i, hi, df⟩, - rw [← (sub_eq_iff_eq_add.mpr df.symm), sub_sub, add_comm, ← sub_sub, sub_self, zero_sub], - refine M.mul_mem_left (-z) (neg_mem_iff.mpr hi) }) (him hz)⟩ + let ⟨y, i, hi, df⟩ := hm.exists_inv hxm, ⟨z, hz⟩ := hx (-y) in + hm.1.1 $ (eq_top_iff_one _).2 $ sub_sub_cancel (z * -y * x + z) 1 ▸ M.sub_mem + (by { rw [mul_assoc, ←mul_add_one, neg_mul, ← (sub_eq_iff_eq_add.mpr df.symm), neg_sub, + sub_add_cancel], + exact M.mul_mem_left _ hi }) (him hz)⟩ lemma exists_mul_sub_mem_of_sub_one_mem_jacobson {I : ideal R} (r : R) - (h : r - 1 ∈ jacobson I) : ∃ s, r * s - 1 ∈ I := + (h : r - 1 ∈ jacobson I) : ∃ s, s * r - 1 ∈ I := begin cases mem_jacobson_iff.1 h 1 with s hs, use s, - simpa [sub_mul] using hs -end - -lemma is_unit_of_sub_one_mem_jacobson_bot (r : R) - (h : r - 1 ∈ jacobson (⊥ : ideal R)) : is_unit r := -begin - cases exists_mul_sub_mem_of_sub_one_mem_jacobson r h with s hs, - rw [mem_bot, sub_eq_zero] at hs, - exact is_unit_of_mul_eq_one _ _ hs + simpa [mul_sub] using hs end /-- An ideal equals its Jacobson radical iff it is the intersection of a set of maximal ideals. @@ -209,9 +194,36 @@ begin refine Inf_le ⟨comap_mono hJ.left, comap_is_maximal_of_surjective _ hf⟩ } end +@[mono] lemma jacobson_mono {I J : ideal R} : I ≤ J → I.jacobson ≤ J.jacobson := +begin + intros h x hx, + erw mem_Inf at ⊢ hx, + exact λ K ⟨hK, hK_max⟩, hx ⟨trans h hK, hK_max⟩ +end + +end ring + +section comm_ring +variables [comm_ring R] [comm_ring S] {I : ideal R} + +lemma radical_le_jacobson : radical I ≤ jacobson I := +le_Inf (λ J hJ, (radical_eq_Inf I).symm ▸ Inf_le ⟨hJ.left, is_maximal.is_prime hJ.right⟩) + +lemma eq_radical_of_eq_jacobson : jacobson I = I → radical I = I := +λ h, le_antisymm (le_trans radical_le_jacobson (le_of_eq h)) le_radical + +lemma is_unit_of_sub_one_mem_jacobson_bot (r : R) + (h : r - 1 ∈ jacobson (⊥ : ideal R)) : is_unit r := +begin + cases exists_mul_sub_mem_of_sub_one_mem_jacobson r h with s hs, + rw [mem_bot, sub_eq_zero, mul_comm] at hs, + exact is_unit_of_mul_eq_one _ _ hs +end + lemma mem_jacobson_bot {x : R} : x ∈ jacobson (⊥ : ideal R) ↔ ∀ y, is_unit (x * y + 1) := ⟨λ hx y, let ⟨z, hz⟩ := (mem_jacobson_iff.1 hx) y in - is_unit_iff_exists_inv.2 ⟨z, by rwa [add_mul, one_mul, ← sub_eq_zero]⟩, + is_unit_iff_exists_inv.2 ⟨z, by rwa [add_mul, one_mul, ← sub_eq_zero, mul_right_comm, + mul_comm _ z, mul_right_comm]⟩, λ h, mem_jacobson_iff.mpr (λ y, (let ⟨b, hb⟩ := is_unit_iff_exists_inv.1 (h y) in ⟨b, (submodule.mem_bot R).2 (hb ▸ (by ring))⟩))⟩ @@ -250,23 +262,20 @@ begin simpa using this } end -@[mono] lemma jacobson_mono {I J : ideal R} : I ≤ J → I.jacobson ≤ J.jacobson := -begin - intros h x hx, - erw mem_Inf at ⊢ hx, - exact λ K ⟨hK, hK_max⟩, hx ⟨trans h hK, hK_max⟩ -end - lemma jacobson_radical_eq_jacobson : I.radical.jacobson = I.jacobson := le_antisymm (le_trans (le_of_eq (congr_arg jacobson (radical_eq_Inf I))) (Inf_le_Inf (λ J hJ, ⟨Inf_le ⟨hJ.1, hJ.2.is_prime⟩, hJ.2⟩))) (jacobson_mono le_radical) +end comm_ring + end jacobson section polynomial open polynomial +variables [comm_ring R] + lemma jacobson_bot_polynomial_le_Inf_map_maximal : jacobson (⊥ : ideal R[X]) ≤ Inf (map C '' {J : ideal R | J.is_maximal}) := begin @@ -297,6 +306,8 @@ end polynomial section is_local +variables [comm_ring R] + /-- An ideal `I` is local iff its Jacobson radical is maximal. -/ class is_local (I : ideal R) : Prop := (out : is_maximal (jacobson I)) @@ -327,7 +338,7 @@ classical.by_cases end is_local -theorem is_primary_of_is_maximal_radical {I : ideal R} (hi : is_maximal (radical I)) : +theorem is_primary_of_is_maximal_radical [comm_ring R] {I : ideal R} (hi : is_maximal (radical I)) : is_primary I := have radical I = jacobson I, from le_antisymm (le_Inf $ λ M ⟨him, hm⟩, hm.is_prime.radical_le_iff.2 him) diff --git a/src/ring_theory/nakayama.lean b/src/ring_theory/nakayama.lean index 6673a59d7c85a..58714ac589368 100644 --- a/src/ring_theory/nakayama.lean +++ b/src/ring_theory/nakayama.lean @@ -52,8 +52,8 @@ begin intros n hn, cases submodule.exists_sub_one_mem_and_smul_eq_zero_of_fg_of_le_smul I N hN hIN with r hr, cases exists_mul_sub_mem_of_sub_one_mem_jacobson r (hIjac hr.1) with s hs, - have : n = (-(r * s - 1) • n), - { rw [neg_sub, sub_smul, mul_comm, mul_smul, hr.2 n hn, one_smul, smul_zero, sub_zero] }, + have : n = (-(s * r - 1) • n), + { rw [neg_sub, sub_smul, mul_smul, hr.2 n hn, one_smul, smul_zero, sub_zero] }, rw this, exact submodule.smul_mem_smul (submodule.neg_mem _ hs) hn end From 9c0dfcd81b268269792f53b748f999f983ab9ba6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Tue, 3 May 2022 20:33:18 +0000 Subject: [PATCH 057/153] doc(order/countable_dense_linear_order): Fix minor mistake (#13921) I wrongfully removed some instances of the word "linear" in #12928. This is in fact used as a hypothesis. --- src/order/countable_dense_linear_order.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/src/order/countable_dense_linear_order.lean b/src/order/countable_dense_linear_order.lean index 969aeef14f134..e2e803c076899 100644 --- a/src/order/countable_dense_linear_order.lean +++ b/src/order/countable_dense_linear_order.lean @@ -10,9 +10,9 @@ import order.ideal ## Results -Suppose `α β` are orders, with `α` countable and `β` dense, nontrivial. Then there is an order -embedding `α ↪ β`. If in addition `α` is dense, nonempty, without endpoints and `β` is countable, -without endpoints, then we can upgrade this to an order isomorphism `α ≃ β`. +Suppose `α β` are linear orders, with `α` countable and `β` dense, nontrivial. Then there is an +order embedding `α ↪ β`. If in addition `α` is dense, nonempty, without endpoints and `β` is +countable, without endpoints, then we can upgrade this to an order isomorphism `α ≃ β`. The idea for both results is to consider "partial isomorphisms", which identify a finite subset of `α` with a finite subset of `β`, and prove that for any such partial isomorphism `f` and `a : α`, we @@ -174,7 +174,7 @@ open partial_iso variables (α β) -/-- Any countable order embeds in any nontrivial dense linear order. -/ +/-- Any countable linear order embeds in any nontrivial dense linear order. -/ theorem embedding_from_countable_to_dense [encodable α] [densely_ordered β] [nontrivial β] : nonempty (α ↪o β) := begin From de07131ef3c3a986eb1ba3c5677d789ff132abe4 Mon Sep 17 00:00:00 2001 From: CumaKoekmen Date: Tue, 3 May 2022 21:57:27 +0000 Subject: [PATCH 058/153] feat(measure_theory/integral/torus_integral): torus integral and its properties (#12892) Define a generalized torus map and prove some basic properties. Define the torus integral and the integrability of functions on a generalized torus, and prove lemmas about them. Co-authored-by: Yury G. Kudryashov --- .../integral/circle_integral.lean | 7 + .../integral/torus_integral.lean | 247 ++++++++++++++++++ 2 files changed, 254 insertions(+) create mode 100644 src/measure_theory/integral/torus_integral.lean diff --git a/src/measure_theory/integral/circle_integral.lean b/src/measure_theory/integral/circle_integral.lean index 78c08e37a4c96..96252bcf4655f 100644 --- a/src/measure_theory/integral/circle_integral.lean +++ b/src/measure_theory/integral/circle_integral.lean @@ -94,6 +94,8 @@ show (coe ⁻¹' ((* I) ⁻¹' (exp ⁻¹' ((*) R ⁻¹' ((+) c ⁻¹' s))))).co circle_map c R θ - c = circle_map 0 R θ := by simp [circle_map] +lemma circle_map_zero (R θ : ℝ) : circle_map 0 R θ = R * exp (θ * I) := zero_add _ + @[simp] lemma abs_circle_map_zero (R : ℝ) (θ : ℝ) : abs (circle_map 0 R θ) = |R| := by simp [circle_map] @@ -278,6 +280,11 @@ def circle_integral (f : ℂ → E) (c : ℂ) (R : ℝ) : E := notation `∮` binders ` in ` `C(` c `, ` R `)` `, ` r:(scoped:60 f, circle_integral f c R) := r +lemma circle_integral_def_Icc (f : ℂ → E) (c : ℂ) (R : ℝ) : + ∮ z in C(c, R), f z = ∫ θ in Icc 0 (2 * π), deriv (circle_map c R) θ • f (circle_map c R θ) := +by simp only [circle_integral, interval_integral.integral_of_le real.two_pi_pos.le, + measure.restrict_congr_set Ioc_ae_eq_Icc] + namespace circle_integral @[simp] lemma integral_radius_zero (f : ℂ → E) (c : ℂ) : ∮ z in C(c, 0), f z = 0 := diff --git a/src/measure_theory/integral/torus_integral.lean b/src/measure_theory/integral/torus_integral.lean new file mode 100644 index 0000000000000..a9ec6d145be3f --- /dev/null +++ b/src/measure_theory/integral/torus_integral.lean @@ -0,0 +1,247 @@ +/- +Copyright (c) 2022 Cuma Kökmen. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Cuma Kökmen, Yury Kudryashov +-/ +import measure_theory.integral.circle_integral + +/-! +# Integral over a torus in `ℂⁿ` + +In this file we define the integral of a function `f : ℂⁿ → E` over a torus +`{z : ℂⁿ | ∀ i, z i ∈ metric.sphere (c i) (R i)}`. In order to do this, we define +`torus_map (c : ℂⁿ) (R θ : ℝⁿ)` to be the point in `ℂⁿ` given by $z_k=c_k+R_ke^{θ_ki}$, +where $i$ is the imaginary unit, then define `torus_integral f c R` as the integral over +the cube $[0, (λ _, 2π)] = \{θ\|∀ k, 0 ≤ θ_k ≤ 2π\}$ of the Jacobian of the +`torus_map` multiplied by `f (torus_map c R θ)`. + +We also define a predicate saying that `f ∘ torus_map c R` is integrable on the cube +`[0, (λ _, 2\pi)]`. + +## Main definitions + +* `torus_map c R`: the generalized multidimensional exponential map from `ℝⁿ` to `ℂⁿ` that sends + $θ=(θ_0,…,θ_{n-1})$ to $z=(z_0,…,z_{n-1})$, where $z_k= c_k + R_ke^{θ_k i}$; + +* `torus_integrable f c R`: a function `f : ℂⁿ → E` is integrable over the generalized torus + with center `c : ℂⁿ` and radius `R : ℝⁿ` if `f ∘ torus_map c R` is integrable on the + closed cube `Icc (0 : ℝⁿ) (λ _, 2 * π)`; + +* `torus_integral f c R`: the integral of a function `f : ℂⁿ → E` over a torus with + center `c ∈ ℂⁿ` and radius `R ∈ ℝⁿ` defined as + $\iiint_{[0, 2 * π]} (∏_{k = 1}^{n} i R_k e^{θ_k * i}) • f (c + Re^{θ_k i})\,dθ_0…dθ_{k-1}$. + +## Main statements + +* `torus_integral_dim0`, `torus_integral_dim1`, `torus_integral_succ`: formulas for `torus_integral` + in cases of dimension `0`, `1`, and `n + 1`. + +## Notations + +- `ℝ⁰`, `ℝ¹`, `ℝⁿ`, `ℝⁿ⁺¹`: local notation for `fin 0 → ℝ`, `fin 1 → ℝ`, `fin n → ℝ`, and + `fin (n + 1) → ℝ`, respectively; +- `ℂ⁰`, `ℂ¹`, `ℂⁿ`, `ℂⁿ⁺¹`: local notation for `fin 0 → ℂ`, `fin 1 → ℂ`, `fin n → ℂ`, and + `fin (n + 1) → ℂ`, respectively; +- `∯ z in T(c, R), f z`: notation for `torus_integral f c R`; +- `∮ z in C(c, R), f z`: notation for `circle_integral f c R`, defined elsewhere; +- `∏ k, f k`: notation for `finset.prod`, defined elsewhere; +- `π`: notation for `real.pi`, defined elsewhere. + +## Tags + +integral, torus +-/ + +variable {n : ℕ} +variables {E : Type*} [normed_group E] + +noncomputable theory + +open complex set measure_theory function filter topological_space +open_locale real big_operators + +local notation `ℝ⁰`:= fin 0 → ℝ +local notation `ℂ⁰`:= fin 0 → ℂ +local notation `ℝ¹`:= fin 1 → ℝ +local notation `ℂ¹`:= fin 1 → ℂ +local notation `ℝⁿ`:= fin n → ℝ +local notation `ℂⁿ`:= fin n → ℂ +local notation `ℝⁿ⁺¹`:= fin (n + 1) → ℝ +local notation `ℂⁿ⁺¹`:= fin (n + 1) → ℂ + +/-! +### `torus_map`, a generalization of a torus +-/ + +/-- The n dimensional exponential map $θ_i ↦ c + R e^{θ_i*I}, θ ∈ ℝⁿ$ representing +a torus in `ℂⁿ` with center `c ∈ ℂⁿ` and generalized radius `R ∈ ℝⁿ`, so we can adjust +it to every n axis. -/ +def torus_map (c : ℂⁿ) (R : ℝⁿ) : ℝⁿ → ℂⁿ := +λ θ i, c i + R i * exp(θ i * I) + +lemma torus_map_sub_center (c : ℂⁿ) (R : ℝⁿ) (θ : ℝⁿ) : + torus_map c R θ - c = torus_map 0 R θ := +by { ext1 i, simp [torus_map] } + +lemma torus_map_eq_center_iff {c : ℂⁿ} {R : ℝⁿ} {θ : ℝⁿ} : + torus_map c R θ = c ↔ R = 0 := +by simp [funext_iff, torus_map, exp_ne_zero] + +@[simp] lemma torus_map_zero_radius (c : ℂⁿ) : torus_map c 0 = const ℝⁿ c := +by { ext1, rw torus_map_eq_center_iff.2 rfl } + +/-! +### Integrability of a function on a generalized torus +-/ + +/-- A function `f : ℂⁿ → E` is integrable on the generalized torus if the function +`f ∘ torus_map c R θ` is integrable on `Icc (0 : ℝⁿ) (λ _, 2 * π)`-/ +def torus_integrable (f : ℂⁿ → E) (c : ℂⁿ) (R : ℝⁿ) : Prop := + integrable_on (λ (θ : ℝⁿ), f (torus_map c R θ)) (Icc (0 : ℝⁿ) (λ _, 2 * π)) volume + +namespace torus_integrable + +variables {f g : ℂⁿ → E} {c : ℂⁿ} {R : ℝⁿ} + +/-- Constant functions are torus integrable -/ +lemma torus_integrable_const (a : E) (c : ℂⁿ) (R : ℝⁿ) : + torus_integrable (λ _, a) c R := +by simp [torus_integrable, measure_Icc_lt_top] + +/-- If `f` is torus integrable then `-f` is torus integrable. -/ +protected lemma neg (hf : torus_integrable f c R) : torus_integrable (-f) c R := hf.neg + +/-- If `f` and `g` are two torus integrable functions, then so is `f + g`. -/ +protected lemma add (hf : torus_integrable f c R) (hg : torus_integrable g c R) : + torus_integrable (f + g) c R := +hf.add hg + +/-- If `f` and `g` are two torus integrable functions, then so is `f - g`. -/ +protected lemma sub (hf : torus_integrable f c R) (hg : torus_integrable g c R) : + torus_integrable (f - g) c R := +hf.sub hg + +lemma torus_integrable_zero_radius {f : ℂⁿ → E} {c : ℂⁿ} : + torus_integrable f c 0 := +begin + rw [torus_integrable, torus_map_zero_radius], + apply torus_integrable_const (f c) c 0, +end + +/--The function given in the definition of `torus_integral` is integrable. -/ +lemma function_integrable [normed_space ℂ E] (hf : torus_integrable f c R) : + integrable_on (λ (θ : ℝⁿ), (∏ i, R i * exp(θ i * I) * I : ℂ) • f (torus_map c R θ)) + (Icc (0 : ℝⁿ) (λ _, 2 * π)) volume := +begin + refine (hf.norm.const_mul (∏ i, |R i|)).mono' _ _, + { refine (continuous.ae_strongly_measurable _).smul hf.1, + exact continuous_finset_prod finset.univ (λ i hi, (continuous_const.mul + (((continuous_of_real.comp (continuous_apply i)).mul continuous_const).cexp)).mul + continuous_const) }, + simp [norm_smul, map_prod], +end + +end torus_integrable + +variables [normed_space ℂ E] [complete_space E] {f g : ℂⁿ → E} {c : ℂⁿ} {R : ℝⁿ} + +/--The definition of the integral over a generalized torus with center `c ∈ ℂⁿ` and radius `R ∈ ℝⁿ` +as the `•`-product of the derivative of `torus_map` and `f (torus_map c R θ)`-/ +def torus_integral (f : ℂⁿ → E) (c : ℂⁿ) (R : ℝⁿ) := +∫ (θ : ℝⁿ) in Icc (0 : ℝⁿ) (λ _, 2 * π), (∏ i, R i * exp(θ i * I) * I : ℂ) • f (torus_map c R θ) + +notation `∯` binders ` in ` `T(` c `, ` R `)` `, ` r:(scoped:60 f, torus_integral f c R) := r + +lemma torus_integral_radius_zero (hn : n ≠ 0) (f : ℂⁿ → E) (c : ℂⁿ): ∯ x in T(c, 0), f x = 0 := +by simp only [torus_integral, pi.zero_apply, of_real_zero, mul_zero, zero_mul, fin.prod_const, + zero_pow' n hn, zero_smul, integral_zero] + +lemma torus_integral_neg (f : ℂⁿ → E) (c : ℂⁿ) (R : ℝⁿ) : + ∯ x in T(c, R), -f x = -∯ x in T(c, R), f x := +by simp [torus_integral, integral_neg] + +lemma torus_integral_add (hf : torus_integrable f c R) (hg : torus_integrable g c R) : + ∯ x in T(c, R), f x + g x = (∯ x in T(c, R), f x) + ∯ x in T(c, R), g x := +by simpa only [torus_integral, smul_add, pi.add_apply] + using integral_add hf.function_integrable hg.function_integrable + +lemma torus_integral_sub (hf : torus_integrable f c R) (hg : torus_integrable g c R) : + ∯ x in T(c, R), f x - g x = (∯ x in T(c, R), f x) - ∯ x in T(c, R), g x := +by simpa only [sub_eq_add_neg, ← torus_integral_neg] using torus_integral_add hf hg.neg + +lemma torus_integral_smul {𝕜 : Type*} [is_R_or_C 𝕜] [normed_space 𝕜 E] [smul_comm_class 𝕜 ℂ E] + (a : 𝕜) (f : ℂⁿ → E) (c : ℂⁿ) (R : ℝⁿ) : + ∯ x in T(c, R), a • f x = a • ∯ x in T(c, R), f x := +by simp only [torus_integral, integral_smul, ← smul_comm a] + +lemma torus_integral_const_mul (a : ℂ) (f : ℂⁿ → ℂ) (c : ℂⁿ) (R : ℝⁿ) : + ∯ x in T(c, R), a * f x = a * ∯ x in T(c, R), f x := +torus_integral_smul a f c R + +/--If for all `θ : ℝⁿ`, `∥f (torus_map c R θ)∥` is less than or equal to a constant `C : ℝ`, then +`∥∯ x in T(c, R), f x∥` is less than or equal to `(2 * π)^n * (∏ i, |R i|) * C`-/ +lemma norm_torus_integral_le_of_norm_le_const {C : ℝ} (hf : ∀ θ, ∥f (torus_map c R θ)∥ ≤ C) : + ∥∯ x in T(c, R), f x∥ ≤ (2 * π)^(n: ℕ) * (∏ i, |R i|) * C := +calc ∥∯ x in T(c, R), f x∥ ≤ (∏ i, |R i|) * C * (volume (Icc (0 : ℝⁿ) (λ _, 2 * π))).to_real : + norm_set_integral_le_of_norm_le_const' measure_Icc_lt_top measurable_set_Icc $ λ θ hθ, + ( calc ∥(∏ i : fin n, R i * exp (θ i * I) * I : ℂ) • f (torus_map c R θ)∥ + = (∏ i : fin n, |R i|) * ∥f (torus_map c R θ)∥ : by simp [norm_smul] + ... ≤ (∏ i : fin n, |R i|) * C : + mul_le_mul_of_nonneg_left (hf _) (finset.prod_nonneg $ λ _ _, abs_nonneg _) ) +... = (2 * π)^(n: ℕ) * (∏ i, |R i|) * C : + by simp only [pi.zero_def, real.volume_Icc_pi_to_real (λ _, real.two_pi_pos.le), sub_zero, + fin.prod_const, mul_assoc, mul_comm ((2 * π) ^ (n : ℕ))] + +@[simp] lemma torus_integral_dim0 (f : ℂ⁰ → E) (c : ℂ⁰) (R : ℝ⁰) : ∯ x in T(c, R), f x = f c := +by simp only [torus_integral, fin.prod_univ_zero, one_smul, + subsingleton.elim (λ i : fin 0, 2 * π) 0, Icc_self, measure.restrict_singleton, volume_pi, + integral_smul_measure, integral_dirac, measure.pi_of_empty _ 0, + measure.dirac_apply_of_mem (mem_singleton _), subsingleton.elim (torus_map c R 0) c] + +/-- In dimension one, `torus_integral` is the same as `circle_integral` +(up to the natural equivalence between `ℂ` and `fin 1 → ℂ`). -/ +lemma torus_integral_dim1 (f : ℂ¹ → E) (c : ℂ¹) (R : ℝ¹) : + ∯ x in T(c, R), f x = ∮ z in C(c 0, R 0), f (λ _, z) := +begin + have : (λ (x : ℝ) (b : fin 1), x) ⁻¹' Icc 0 (λ _, 2 * π) = Icc 0 (2 * π), + from (order_iso.fun_unique (fin 1) ℝ).symm.preimage_Icc _ _, + simp only [torus_integral, circle_integral, interval_integral.integral_of_le real.two_pi_pos.le, + measure.restrict_congr_set Ioc_ae_eq_Icc, deriv_circle_map, fin.prod_univ_one, + ← (volume_preserving_fun_unique (fin 1) ℝ).symm.set_integral_preimage_emb + (measurable_equiv.measurable_embedding _), this, measurable_equiv.fun_unique_symm_apply], + simp only [torus_map, circle_map, zero_add], + rcongr +end + +/-- Recurrent formula for `torus_integral`, see also `torus_integral_succ`. -/ +lemma torus_integral_succ_above {f : ℂⁿ⁺¹ → E} {c : ℂⁿ⁺¹} {R : ℝⁿ⁺¹} (hf : torus_integrable f c R) + (i : fin (n + 1)) : + ∯ x in T(c, R), f x = + ∮ x in C(c i, R i), ∯ y in T(c ∘ i.succ_above, R ∘ i.succ_above), f (i.insert_nth x y) := +begin + set e : ℝ × ℝⁿ ≃ᵐ ℝⁿ⁺¹ := (measurable_equiv.pi_fin_succ_above_equiv (λ _, ℝ) i).symm, + have hem : measure_preserving e, + from (volume_preserving_pi_fin_succ_above_equiv (λ j : fin (n + 1), ℝ) i).symm, + have heπ : e ⁻¹' (Icc 0 (λ _, 2 * π)) = Icc 0 (2 * π) ×ˢ Icc (0 : ℝⁿ) (λ _, 2 * π), + from ((order_iso.pi_fin_succ_above_iso (λ _, ℝ) i).symm.preimage_Icc _ _).trans + (Icc_prod_eq _ _), + rw [torus_integral, ← hem.map_eq, set_integral_map_equiv, heπ, measure.volume_eq_prod, + set_integral_prod, circle_integral_def_Icc], + { refine set_integral_congr measurable_set_Icc (λ θ hθ, _), + simp only [torus_integral, ← integral_smul, deriv_circle_map, i.prod_univ_succ_above _, + smul_smul, torus_map, circle_map_zero], + refine set_integral_congr measurable_set_Icc (λ Θ hΘ, _), + simp only [measurable_equiv.pi_fin_succ_above_equiv_symm_apply, i.insert_nth_apply_same, + i.insert_nth_apply_succ_above, (∘)], + congr' 2, + simp only [funext_iff, i.forall_iff_succ_above, circle_map, fin.insert_nth_apply_same, + eq_self_iff_true, fin.insert_nth_apply_succ_above, implies_true_iff, and_self] }, + { have := hf.function_integrable, + rwa [← hem.integrable_on_comp_preimage e.measurable_embedding, heπ] at this } +end + +/-- Recurrent formula for `torus_integral`, see also `torus_integral_succ_above`. -/ +lemma torus_integral_succ {f : ℂⁿ⁺¹ → E} {c : ℂⁿ⁺¹} {R : ℝⁿ⁺¹} (hf : torus_integrable f c R) : + ∯ x in T(c, R), f x = + ∮ x in C(c 0, R 0), ∯ y in T(c ∘ fin.succ, R ∘ fin.succ), f (fin.cons x y) := +by simpa using torus_integral_succ_above hf 0 From fd65159725ca63778da71bcd25b2190489a979d6 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 3 May 2022 22:34:05 +0000 Subject: [PATCH 059/153] feat(topology/metric_space/basic): golf, avoid unfold (#13923) * Don't use `unfold` in `nnreal.pseudo_metric_space`. * Golf some proofs. --- src/topology/metric_space/basic.lean | 47 ++++++++-------------------- 1 file changed, 13 insertions(+), 34 deletions(-) diff --git a/src/topology/metric_space/basic.lean b/src/topology/metric_space/basic.lean index 61eea9f05b9bc..35a6e65f133cd 100644 --- a/src/topology/metric_space/basic.lean +++ b/src/topology/metric_space/basic.lean @@ -1086,17 +1086,8 @@ theorem metric.complete_of_convergent_controlled_sequences (B : ℕ → real) (h (H : ∀u : ℕ → α, (∀N n m : ℕ, N ≤ n → N ≤ m → dist (u n) (u m) < B N) → ∃x, tendsto u at_top (𝓝 x)) : complete_space α := -begin - -- this follows from the same criterion in emetric spaces. We just need to translate - -- the convergence assumption from `dist` to `edist` - apply emetric.complete_of_convergent_controlled_sequences (λn, ennreal.of_real (B n)), - { simp [hB] }, - { assume u Hu, - apply H, - assume N n m hn hm, - rw [← ennreal.of_real_lt_of_real_iff (hB N), ← edist_dist], - exact Hu N n m hn hm } -end +uniform_space.complete_of_convergent_controlled_sequences + (λ n, {p:α×α | dist p.1 p.2 < B n}) (λ n, dist_mem_uniformity $ hB n) H theorem metric.complete_of_cauchy_seq_tendsto : (∀ u : ℕ → α, cauchy_seq u → ∃a, tendsto u at_top (𝓝 a)) → complete_space α := @@ -1366,20 +1357,19 @@ end mul_opposite section nnreal -noncomputable instance : pseudo_metric_space ℝ≥0 := by unfold nnreal; apply_instance +noncomputable instance : pseudo_metric_space ℝ≥0 := subtype.pseudo_metric_space lemma nnreal.dist_eq (a b : ℝ≥0) : dist a b = |(a:ℝ) - b| := rfl lemma nnreal.nndist_eq (a b : ℝ≥0) : nndist a b = max (a - b) (b - a) := begin - wlog h : a ≤ b, - { apply nnreal.coe_eq.1, - rw [tsub_eq_zero_iff_le.2 h, max_eq_right (zero_le $ b - a), ← dist_nndist, nnreal.dist_eq, - nnreal.coe_sub h, abs_eq_max_neg, neg_sub], - apply max_eq_right, - linarith [nnreal.coe_le_coe.2 h] }, - rwa [nndist_comm, max_comm] + /- WLOG, `b ≤ a`. `wlog h : b ≤ a` works too but it is much slower because Lean tries to prove one + case from the other and fails; `tactic.skip` tells Lean not to try. -/ + wlog h : b ≤ a := le_total b a using [a b, b a] tactic.skip, + { rw [← nnreal.coe_eq, ← dist_nndist, nnreal.dist_eq, tsub_eq_zero_iff_le.2 h, + max_eq_left (zero_le $ a - b), ← nnreal.coe_sub h, abs_of_nonneg (a - b).coe_nonneg] }, + { rwa [nndist_comm, max_comm] } end @[simp] lemma nnreal.nndist_zero_eq_val (z : ℝ≥0) : nndist 0 z = z := @@ -1544,23 +1534,12 @@ subset.antisymm lemma dense_iff {s : set α} : dense s ↔ ∀ x, ∀ r > 0, (ball x r ∩ s).nonempty := -begin - apply forall_congr (λ x, _), - rw mem_closure_iff, - refine forall_congr (λ ε, forall_congr (λ h, exists_congr (λ y, _))), - rw [mem_inter_iff, mem_ball', exists_prop, and_comm] -end +forall_congr $ λ x, by simp only [mem_closure_iff, set.nonempty, exists_prop, mem_inter_eq, + mem_ball', and_comm] lemma dense_range_iff {f : β → α} : dense_range f ↔ ∀ x, ∀ r > 0, ∃ y, dist x (f y) < r := -begin - rw [dense_range, metric.dense_iff], - refine forall_congr (λ x, forall_congr (λ r, forall_congr (λ rpos, ⟨_, _⟩))), - { rintros ⟨-, hz, ⟨z, rfl⟩⟩, - exact ⟨z, metric.mem_ball'.1 hz⟩ }, - { rintros ⟨z, hz⟩, - exact ⟨f z, metric.mem_ball'.1 hz, mem_range_self _⟩ } -end +forall_congr $ λ x, by simp only [mem_closure_iff, exists_range_iff] /-- If a set `s` is separable, then the corresponding subtype is separable in a metric space. This is not obvious, as the countable set whose closure covers `s` does not need in general to @@ -1651,7 +1630,7 @@ begin end lemma nndist_pi_def (f g : Πb, π b) : nndist f g = sup univ (λb, nndist (f b) (g b)) := -subtype.eta _ _ +nnreal.eq rfl lemma dist_pi_def (f g : Πb, π b) : dist f g = (sup univ (λb, nndist (f b) (g b)) : ℝ≥0) := rfl From 6c0580a2e81c8320725548e3f18b6a299264b9a2 Mon Sep 17 00:00:00 2001 From: Bryan Gin-ge Chen Date: Tue, 3 May 2022 23:37:38 +0000 Subject: [PATCH 060/153] fix(.docker/*): update elan URL (#13928) These are hopefully the last occurrences of the URL that was breaking things earlier today. cf. #13906 --- .docker/debian/lean/Dockerfile | 2 +- .docker/gitpod/mathlib/Dockerfile | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/.docker/debian/lean/Dockerfile b/.docker/debian/lean/Dockerfile index d06a6ed7f6a73..e43de629063d2 100644 --- a/.docker/debian/lean/Dockerfile +++ b/.docker/debian/lean/Dockerfile @@ -49,7 +49,7 @@ ENTRYPOINT ["/bin/bash", "-l"] ENV PATH="/home/lean/.elan/bin:/home/lean/.local/bin:$PATH" # install elan -RUN curl https://raw.githubusercontent.com/Kha/elan/master/elan-init.sh -sSf | sh -s -- -y --default-toolchain none && \ +RUN curl https://raw.githubusercontent.com/leanprover/elan/master/elan-init.sh -sSf | sh -s -- -y --default-toolchain none && \ . ~/.profile && \ elan toolchain install $(curl https://raw.githubusercontent.com/leanprover-community/mathlib/master/leanpkg.toml | grep lean_version | awk -F'"' '{print $2}') # install leanproject diff --git a/.docker/gitpod/mathlib/Dockerfile b/.docker/gitpod/mathlib/Dockerfile index 234b6100f0f5f..f48dc06c22c8e 100644 --- a/.docker/gitpod/mathlib/Dockerfile +++ b/.docker/gitpod/mathlib/Dockerfile @@ -14,7 +14,7 @@ WORKDIR /home/gitpod SHELL ["/bin/bash", "-c"] # install elan -RUN curl https://raw.githubusercontent.com/Kha/elan/master/elan-init.sh -sSf | sh -s -- -y --default-toolchain none +RUN curl https://raw.githubusercontent.com/leanprover/elan/master/elan-init.sh -sSf | sh -s -- -y --default-toolchain none # install whichever toolchain mathlib is currently using RUN . ~/.profile && elan toolchain install $(curl https://raw.githubusercontent.com/leanprover-community/mathlib/master/leanpkg.toml | grep lean_version | awk -F'"' '{print $2}') # install `leanproject` using `pip` From fc3de1972fd9f6f215b9d3910a093b3dec92fc51 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Wed, 4 May 2022 00:13:27 +0000 Subject: [PATCH 061/153] feat(ring_theory/ideal/local_ring): generalize lemmas to semirings (#13471) What is essentially new is the proof of `local_ring.of_surjective` and `local_ring.is_unit_or_is_unit_of_is_unit_add`. - I changed the definition of local ring to `local_ring.of_is_unit_or_is_unit_of_add_one`, which is reminiscent of the definition before the recent change in #13341. The equivalence of the previous definition is essentially given by `local_ring.is_unit_or_is_unit_of_is_unit_add`. The choice of the definition is insignificant here because they are all equivalent, but I think the choice here is better for the default constructor because this condition is "weaker" than e.g. `local_ring.of_non_units_add` in some sense. - The proof of `local_ring.of_surjective` needs `[is_local_ring_hom f]`, which was not necessary for commutative rings in the previous proof. So the new version here is not a genuine generalization of the previous version. The previous version was renamed to `local_ring.of_surjective'`. --- src/logic/equiv/transfer_instance.lean | 6 +- src/number_theory/padics/padic_integers.lean | 2 +- src/ring_theory/ideal/local_ring.lean | 87 +++++++++++++------- src/ring_theory/localization/at_prime.lean | 2 +- 4 files changed, 60 insertions(+), 37 deletions(-) diff --git a/src/logic/equiv/transfer_instance.lean b/src/logic/equiv/transfer_instance.lean index 60ed443084927..e6ec8b2e8628c 100644 --- a/src/logic/equiv/transfer_instance.lean +++ b/src/logic/equiv/transfer_instance.lean @@ -394,11 +394,11 @@ end equiv namespace ring_equiv -protected lemma local_ring {A B : Type*} [comm_ring A] [local_ring A] [comm_ring B] (e : A ≃+* B) : - local_ring B := +protected lemma local_ring {A B : Type*} [comm_semiring A] [local_ring A] [comm_semiring B] + (e : A ≃+* B) : local_ring B := begin haveI := e.symm.to_equiv.nontrivial, - refine @local_ring.of_surjective A B _ _ _ _ e e.to_equiv.surjective, + exact local_ring.of_surjective (e : A →+* B) e.surjective end end ring_equiv diff --git a/src/number_theory/padics/padic_integers.lean b/src/number_theory/padics/padic_integers.lean index c33bd945866fc..509275061ac22 100644 --- a/src/number_theory/padics/padic_integers.lean +++ b/src/number_theory/padics/padic_integers.lean @@ -537,7 +537,7 @@ section dvr /-! ### Discrete valuation ring -/ instance : local_ring ℤ_[p] := -local_ring.mk $ by simp only [mem_nonunits]; exact λ x y, norm_lt_one_add +local_ring.of_nonunits_add $ by simp only [mem_nonunits]; exact λ x y, norm_lt_one_add lemma p_nonnunit : (p : ℤ_[p]) ∈ nonunits ℤ_[p] := have (p : ℝ)⁻¹ < 1, from inv_lt_one $ by exact_mod_cast hp_prime.1.one_lt, diff --git a/src/ring_theory/ideal/local_ring.lean b/src/ring_theory/ideal/local_ring.lean index 468799815127e..77394649e09d7 100644 --- a/src/ring_theory/ideal/local_ring.lean +++ b/src/ring_theory/ideal/local_ring.lean @@ -16,9 +16,9 @@ Define local rings as commutative rings having a unique maximal ideal. ## Main definitions -* `local_ring`: A predicate on commutative semirings, stating that the set of nonunits is closed - under the addition. This is shown to be equivalent to the condition that there exists a unique - maximal ideal. +* `local_ring`: A predicate on commutative semirings, stating that for any pair of elements that + adds up to `1`, one of them is a unit. This is shown to be equivalent to the condition that there + exists a unique maximal ideal. * `local_ring.maximal_ideal`: The unique maximal ideal for a local rings. Its carrier set is the set of non units. * `is_local_ring_hom`: A predicate on semiring homomorphisms, requiring that it maps nonunits @@ -32,27 +32,39 @@ universes u v w u' variables {R : Type u} {S : Type v} {T : Type w} {K : Type u'} -/-- A semiring is local if it is nontrivial and the set of nonunits is closed under the addition. +/-- A semiring is local if it is nontrivial and `a` or `b` is a unit whenever `a + b = 1`. Note that `local_ring` is a predicate. -/ class local_ring (R : Type u) [semiring R] extends nontrivial R : Prop := -(nonunits_add : ∀ {a b : R}, a ∈ nonunits R → b ∈ nonunits R → a + b ∈ nonunits R) +of_is_unit_or_is_unit_of_add_one :: +(is_unit_or_is_unit_of_add_one {a b : R} (h : a + b = 1) : is_unit a ∨ is_unit b) section comm_semiring variables [comm_semiring R] namespace local_ring +lemma of_is_unit_or_is_unit_of_is_unit_add [nontrivial R] + (h : ∀ a b : R, is_unit (a + b) → is_unit a ∨ is_unit b) : + local_ring R := +⟨λ a b hab, h a b $ hab.symm ▸ is_unit_one⟩ + +/-- A semiring is local if it is nontrivial and the set of nonunits is closed under the addition. -/ +lemma of_nonunits_add [nontrivial R] + (h : ∀ a b : R, a ∈ nonunits R → b ∈ nonunits R → a + b ∈ nonunits R) : + local_ring R := +⟨λ a b hab, or_iff_not_and_not.2 $ λ H, h a b H.1 H.2 $ hab.symm ▸ is_unit_one⟩ + /-- A semiring is local if it has a unique maximal ideal. -/ lemma of_unique_max_ideal (h : ∃! I : ideal R, I.is_maximal) : local_ring R := -@local_ring.mk _ _ (nontrivial_of_ne (0 : R) 1 $ +@of_nonunits_add _ _ (nontrivial_of_ne (0 : R) 1 $ let ⟨I, Imax, _⟩ := h in (λ (H : 0 = 1), Imax.1.1 $ I.eq_top_iff_one.2 $ H ▸ I.zero_mem)) $ λ x y hx hy H, let ⟨I, Imax, Iuniq⟩ := h in let ⟨Ix, Ixmax, Hx⟩ := exists_max_ideal_of_mem_nonunits hx in let ⟨Iy, Iymax, Hy⟩ := exists_max_ideal_of_mem_nonunits hy in - have xmemI : x ∈ I, from ((Iuniq Ix Ixmax) ▸ Hx), - have ymemI : y ∈ I, from ((Iuniq Iy Iymax) ▸ Hy), + have xmemI : x ∈ I, from Iuniq Ix Ixmax ▸ Hx, + have ymemI : y ∈ I, from Iuniq Iy Iymax ▸ Hy, Imax.1.1 $ I.eq_top_of_is_unit_mem (I.add_mem xmemI ymemI) H lemma of_unique_nonzero_prime (h : ∃! P : ideal R, P ≠ ⊥ ∧ ideal.is_prime P) : @@ -66,7 +78,21 @@ of_unique_max_ideal begin exact hPnot_top (hM.1.2 P (bot_lt_iff_ne_bot.2 hPnonzero)) }, end -variables (R) [local_ring R] +variables [local_ring R] + +lemma is_unit_or_is_unit_of_is_unit_add {a b : R} (h : is_unit (a + b)) : + is_unit a ∨ is_unit b := +begin + rcases h with ⟨u, hu⟩, + replace hu : ↑u⁻¹ * a + ↑u⁻¹ * b = 1, from by rw [←mul_add, ←hu, units.inv_mul], + cases is_unit_or_is_unit_of_add_one hu; [left, right]; + exact (is_unit_of_mul_is_unit_right (by assumption)) +end + +lemma nonunits_add {a b : R} (ha : a ∈ nonunits R) (hb : b ∈ nonunits R) : a + b ∈ nonunits R:= +λ H, not_or ha hb (is_unit_or_is_unit_of_is_unit_add H) + +variables (R) /-- The ideal of elements that are not units. -/ def maximal_ideal : ideal R := @@ -115,28 +141,12 @@ namespace local_ring lemma of_is_unit_or_is_unit_one_sub_self [nontrivial R] (h : ∀ a : R, is_unit a ∨ is_unit (1 - a)) : local_ring R := -local_ring.mk -begin - intros x y hx hy, - rintros ⟨u, hu⟩, - apply hy, - suffices : is_unit ((↑u⁻¹ : R) * y), - { rcases this with ⟨s, hs⟩, - use u * s, - convert congr_arg (λ z, (u : R) * z) hs, - rw ← mul_assoc, simp }, - rw show (↑u⁻¹ * y) = (1 - ↑u⁻¹ * x), - { rw eq_sub_iff_add_eq, - replace hu := congr_arg (λ z, (↑u⁻¹ : R) * z) hu.symm, - simpa [mul_add, add_comm] using hu }, - apply or_iff_not_imp_left.1 (h _), - exact mul_mem_nonunits_right hx -end +⟨λ a b hab, add_sub_cancel' a b ▸ hab.symm ▸ h a⟩ variables [local_ring R] lemma is_unit_or_is_unit_one_sub_self (a : R) : is_unit a ∨ is_unit (1 - a) := -or_iff_not_and_not.2 $ λ h, nonunits_add h.1 h.2 $ (add_sub_cancel'_right a 1).symm ▸ is_unit_one +is_unit_or_is_unit_of_is_unit_add $ (add_sub_cancel'_right a 1).symm ▸ is_unit_one lemma is_unit_of_mem_nonunits_one_sub_self (a : R) (h : 1 - a ∈ nonunits R) : is_unit a := @@ -146,7 +156,7 @@ lemma is_unit_one_sub_self_of_mem_nonunits (a : R) (h : a ∈ nonunits R) : is_unit (1 - a) := or_iff_not_imp_left.1 (is_unit_or_is_unit_one_sub_self a) h -lemma of_surjective [comm_ring S] [nontrivial S] (f : R →+* S) (hf : function.surjective f) : +lemma of_surjective' [comm_ring S] [nontrivial S] (f : R →+* S) (hf : function.surjective f) : local_ring S := of_is_unit_or_is_unit_one_sub_self begin @@ -187,11 +197,11 @@ instance is_local_ring_hom_comp { map_nonunit := λ a, is_local_ring_hom.map_nonunit a ∘ is_local_ring_hom.map_nonunit (f a) } instance is_local_ring_hom_equiv (f : R ≃+* S) : - is_local_ring_hom f.to_ring_hom := + is_local_ring_hom (f : R →+* S) := { map_nonunit := λ a ha, begin - convert f.symm.to_ring_hom.is_unit_map ha, - rw ring_equiv.symm_to_ring_hom_apply_to_ring_hom_apply, + convert (f.symm : S →+* R).is_unit_map ha, + exact (ring_equiv.symm_apply_apply f a).symm, end } @[simp] lemma is_unit_of_map_unit (f : R →+* S) [is_local_ring_hom f] @@ -217,7 +227,7 @@ lemma _root_.ring_hom.domain_local_ring {R S : Type*} [comm_semiring R] [comm_se [is_local_ring_hom f] : _root_.local_ring R := begin haveI : nontrivial R := pullback_nonzero f f.map_zero f.map_one, - constructor, + apply local_ring.of_nonunits_add, intros a b, simp_rw [←map_mem_nonunits_iff f, f.map_add], exact local_ring.nonunits_add @@ -284,6 +294,19 @@ end end +lemma of_surjective [comm_semiring R] [local_ring R] [comm_semiring S] [nontrivial S] + (f : R →+* S) [is_local_ring_hom f] (hf : function.surjective f) : + local_ring S := +of_is_unit_or_is_unit_of_is_unit_add +begin + intros a b hab, + obtain ⟨a, rfl⟩ := hf a, + obtain ⟨b, rfl⟩ := hf b, + replace hab : is_unit (f (a + b)), from by simpa only [map_add] using hab, + exact (is_unit_or_is_unit_of_is_unit_add $ is_local_ring_hom.map_nonunit _ hab).imp + f.is_unit_map f.is_unit_map +end + section variables (R) [comm_ring R] [local_ring R] [comm_ring S] [local_ring S] diff --git a/src/ring_theory/localization/at_prime.lean b/src/ring_theory/localization/at_prime.lean index 6a7b6b328df44..b628e3e55cb0f 100644 --- a/src/ring_theory/localization/at_prime.lean +++ b/src/ring_theory/localization/at_prime.lean @@ -59,7 +59,7 @@ protected abbreviation localization.at_prime := localization I.prime_compl namespace is_localization theorem at_prime.local_ring [is_localization.at_prime S I] : local_ring S := -@local_ring.mk _ _ (nontrivial_of_ne (0 : S) 1 +@local_ring.of_nonunits_add _ _ (nontrivial_of_ne (0 : S) 1 (λ hze, begin rw [←(algebra_map R S).map_one, ←(algebra_map R S).map_zero] at hze, obtain ⟨t, ht⟩ := (eq_iff_exists I.prime_compl S).1 hze, From dd584384a10448029c50519a25803c85b22789ed Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 4 May 2022 01:08:00 +0000 Subject: [PATCH 062/153] feat(set_theory/game/short): Birthday of short games (#13875) We prove that a short game has a finite birthday. We also clean up the file somewhat. --- src/set_theory/game/short.lean | 54 ++++++++++++++++++---------------- 1 file changed, 28 insertions(+), 26 deletions(-) diff --git a/src/set_theory/game/short.lean b/src/set_theory/game/short.lean index 92b9c76a0c801..83bacb45a358e 100644 --- a/src/set_theory/game/short.lean +++ b/src/set_theory/game/short.lean @@ -4,7 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ import data.fintype.basic +import set_theory.cardinal.cofinality import set_theory.game.basic +import set_theory.game.birthday /-! # Short games @@ -90,11 +92,28 @@ def move_right_short' {xl xr} (xL xR) [S : short (mk xl xr xL xR)] (j : xr) : sh by { casesI S with _ _ _ _ _ R _ _, apply R } local attribute [instance] move_right_short' -instance short.of_pempty {xL} {xR} : short (mk pempty pempty xL xR) := -short.mk (λ i, pempty.elim i) (λ j, pempty.elim j) +theorem short_birthday : ∀ (x : pgame.{u}) [short x], x.birthday < ordinal.omega +| ⟨xl, xr, xL, xR⟩ hs := +begin + haveI := hs, + unfreezingI { rcases hs with ⟨_, _, _, _, sL, sR, hl, hr⟩ }, + rw [birthday, max_lt_iff], + split, all_goals + { rw ←cardinal.ord_omega, + refine cardinal.lsub_lt_ord_of_is_regular.{u u} cardinal.is_regular_omega + (cardinal.lt_omega_of_fintype _) (λ i, _), + rw cardinal.ord_omega, + apply short_birthday _ }, + { exact move_left_short' xL xR i }, + { exact move_right_short' xL xR i } +end + +/-- This leads to infinite loops if made into an instance. -/ +def short.of_is_empty {l r xL xR} [is_empty l] [is_empty r] : short (mk l r xL xR) := +short.mk is_empty_elim is_empty_elim instance short_0 : short 0 := -short.mk (λ i, by cases i) (λ j, by cases j) +short.of_is_empty instance short_1 : short 1 := short.mk (λ i, begin cases i, apply_instance, end) (λ j, by cases j) @@ -125,43 +144,26 @@ def short_of_relabelling : Π {x y : pgame.{u}} (R : relabelling x y) (S : short | x y ⟨L, R, rL, rR⟩ S := begin resetI, - haveI := (fintype.of_equiv _ L), - haveI := (fintype.of_equiv _ R), + haveI := fintype.of_equiv _ L, + haveI := fintype.of_equiv _ R, exact short.mk' (λ i, by { rw ←(L.right_inv i), apply short_of_relabelling (rL (L.symm i)) infer_instance, }) (λ j, short_of_relabelling (rR j) infer_instance) end -/-- If `x` has no left move or right moves, it is (very!) short. -/ -def short_of_equiv_empty {x : pgame.{u}} - (el : x.left_moves ≃ pempty) (er : x.right_moves ≃ pempty) : short x := -short_of_relabelling (relabel_relabelling el er).symm short.of_pempty - instance short_neg : Π (x : pgame.{u}) [short x], short (-x) | (mk xl xr xL xR) _ := -begin - resetI, - apply short.mk, - { rintro i, - apply short_neg _, - apply_instance, }, - { rintro j, - apply short_neg _, - apply_instance, } -end +by { resetI, exact short.mk (λ i, short_neg _) (λ i, short_neg _) } using_well_founded { dec_tac := pgame_wf_tac } instance short_add : Π (x y : pgame.{u}) [short x] [short y], short (x + y) | (mk xl xr xL xR) (mk yl yr yL yR) _ _ := begin resetI, - apply short.mk, + apply short.mk, all_goals { rintro ⟨i⟩, - { apply short_add, }, - { change short (mk xl xr xL xR + yL i), apply short_add, } }, - { rintro ⟨j⟩, - { apply short_add, }, - { change short (mk xl xr xL xR + yR j), apply short_add, } }, + { apply short_add }, + { change short (mk xl xr xL xR + _), apply short_add } } end using_well_founded { dec_tac := pgame_wf_tac } From a2a873fc0d70e8488ddabd82194ecfa7e03015e9 Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Wed, 4 May 2022 04:20:31 +0000 Subject: [PATCH 063/153] chore(scripts): update nolints.txt (#13932) I am happy to remove some nolints for you! --- scripts/nolints.txt | 2 -- 1 file changed, 2 deletions(-) diff --git a/scripts/nolints.txt b/scripts/nolints.txt index 1a3a7e67505f9..ccb5f9922f946 100644 --- a/scripts/nolints.txt +++ b/scripts/nolints.txt @@ -1084,8 +1084,6 @@ apply_nolint nonarchimedean_group.prod_self_subset to_additive_doc apply_nolint nonarchimedean_group.prod_subset to_additive_doc -- topology/algebra/order/compact.lean -apply_nolint continuous.bdd_above_range_of_has_compact_mul_support to_additive_doc -apply_nolint continuous.bdd_below_range_of_has_compact_mul_support to_additive_doc apply_nolint continuous.exists_forall_ge_of_has_compact_mul_support to_additive_doc apply_nolint continuous.exists_forall_le_of_has_compact_mul_support to_additive_doc From c1f329df4f29d9411585df27b38dcfc180ff55ae Mon Sep 17 00:00:00 2001 From: tb65536 Date: Wed, 4 May 2022 05:16:36 +0000 Subject: [PATCH 064/153] =?UTF-8?q?feat(data/zmod/quotient):=20The=20quoti?= =?UTF-8?q?ent=20`/stab(b)`=20is=20cyclic=20of=20order=20`minimal=5Fper?= =?UTF-8?q?iod=20((+=E1=B5=A5)=20a)=20b`=20(#13722)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds an isomorphism stating that the quotient `/stab(b)` is cyclic of order `minimal_period ((+ᵥ) a) b`. There is also a multiplicative version, but it is easily proved from the additive version, so I'll PR the multiplicative version afterwards. Co-authored-by: Scott Morrison --- src/data/zmod/quotient.lean | 28 ++++++++++++++++++++++++++++ 1 file changed, 28 insertions(+) diff --git a/src/data/zmod/quotient.lean b/src/data/zmod/quotient.lean index 3fd7a8587f684..da2352fce381c 100644 --- a/src/data/zmod/quotient.lean +++ b/src/data/zmod/quotient.lean @@ -63,3 +63,31 @@ def quotient_span_equiv_zmod (a : ℤ) : (quotient_span_nat_equiv_zmod a.nat_abs) end int + +namespace add_action + +open add_subgroup add_monoid_hom add_equiv function + +variables {α β : Type*} [add_group α] (a : α) [add_action α β] (b : β) + +/-- The quotient `(ℤ ∙ a) ⧸ (stabilizer b)` is cyclic of order `minimal_period ((+ᵥ) a) b`. -/ +noncomputable def zmultiples_quotient_stabilizer_equiv : + zmultiples a ⧸ stabilizer (zmultiples a) b ≃+ zmod (minimal_period ((+ᵥ) a) b) := +(of_bijective (map _ (stabilizer (zmultiples a) b) + (zmultiples_hom (zmultiples a) ⟨a, mem_zmultiples a⟩) (by + { rw [zmultiples_le, mem_comap, mem_stabilizer_iff, + zmultiples_hom_apply, coe_nat_zsmul, ←vadd_iterate], + exact is_periodic_pt_minimal_period ((+ᵥ) a) b })) ⟨by + { rw [←ker_eq_bot_iff, eq_bot_iff], + refine λ q, induction_on' q (λ n hn, _), + rw [mem_bot, eq_zero_iff, int.mem_zmultiples_iff, ←zsmul_vadd_eq_iff_minimal_period_dvd], + exact (eq_zero_iff _).mp hn }, + λ q, induction_on' q (λ ⟨_, n, rfl⟩, ⟨n, rfl⟩)⟩).symm.trans + (int.quotient_zmultiples_nat_equiv_zmod (minimal_period ((+ᵥ) a) b)) + +lemma zmultiples_quotient_stabilizer_equiv_symm_apply (n : zmod (minimal_period ((+ᵥ) a) b)) : + (zmultiples_quotient_stabilizer_equiv a b).symm n = + (n : ℤ) • (⟨a, mem_zmultiples a⟩ : zmultiples a) := +rfl + +end add_action From 4d0b6301ef962e6bd74462998943ce38b7adb132 Mon Sep 17 00:00:00 2001 From: Yuma Mizuno Date: Wed, 4 May 2022 05:51:33 +0000 Subject: [PATCH 065/153] feat(category_theory/bicategory/coherence_tactic): coherence tactic for bicategories (#13417) This PR extends the coherence tactic for monoidal categories #13125 to bicategories. The setup is the same as for monoidal case except for the following : we normalize 2-morphisms before running the coherence tactic. This normalization is achieved by the set of simp lemmas in `whisker_simps` defined in `coherence_tactic.lean`. As a test of the tactic in the real world, I have proved several properties of adjunction in bicategories in #13418. Unfortunately some proofs cause timeout, so it seems that we need to speed up the coherence tactic in the future. Co-authored-by: Scott Morrison --- .../bicategory/coherence_tactic.lean | 244 ++++++++++++++++++ src/category_theory/monoidal/coherence.lean | 47 ++-- .../monoidal/internal/types.lean | 2 + test/coherence.lean | 68 ++--- 4 files changed, 310 insertions(+), 51 deletions(-) create mode 100644 src/category_theory/bicategory/coherence_tactic.lean diff --git a/src/category_theory/bicategory/coherence_tactic.lean b/src/category_theory/bicategory/coherence_tactic.lean new file mode 100644 index 0000000000000..7b65cfecc8849 --- /dev/null +++ b/src/category_theory/bicategory/coherence_tactic.lean @@ -0,0 +1,244 @@ +/- +Copyright (c) 2022 Yuma Mizuno. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yuma Mizuno +-/ +import category_theory.bicategory.coherence + +/-! +# A `coherence` tactic for bicategories, and `⊗≫` (composition up to associators) + +We provide a `coherence` tactic, +which proves that any two 2-morphisms (with the same source and target) +in a bicategory which are built out of associators and unitors +are equal. + +We also provide `f ⊗≫ g`, the `bicategorical_comp` operation, +which automatically inserts associators and unitors as needed +to make the target of `f` match the source of `g`. + +This file mainly deals with the type class setup for the coherence tactic. The actual front end +tactic is given in `category_theory/monooidal/coherence.lean` at the same time as the coherence +tactic for monoidal categories. +-/ + +noncomputable theory + +universes w v u + +open category_theory +open category_theory.free_bicategory +open_locale bicategory + +variables {B : Type u} [bicategory.{w v} B] {a b c d e : B} + +namespace category_theory.bicategory + +/-- A typeclass carrying a choice of lift of a 1-morphism from `B` to `free_bicategory B`. -/ +class lift_hom {a b : B} (f : a ⟶ b) := +(lift : of.obj a ⟶ of.obj b) + +instance lift_hom_id : lift_hom (𝟙 a) := { lift := 𝟙 (of.obj a), } +instance lift_hom_comp (f : a ⟶ b) (g : b ⟶ c) [lift_hom f] [lift_hom g] : lift_hom (f ≫ g) := +{ lift := lift_hom.lift f ≫ lift_hom.lift g, } +@[priority 100] +instance lift_hom_of (f : a ⟶ b) : lift_hom f := { lift := of.map f, } + +/-- A typeclass carrying a choice of lift of a 2-morphism from `B` to `free_bicategory B`. -/ +class lift_hom₂ {f g : a ⟶ b} [lift_hom f] [lift_hom g] (η : f ⟶ g) := +(lift : lift_hom.lift f ⟶ lift_hom.lift g) + +instance lift_hom₂_id (f : a ⟶ b) [lift_hom f] : lift_hom₂ (𝟙 f) := +{ lift := 𝟙 _, } +instance lift_hom₂_left_unitor_hom (f : a ⟶ b) [lift_hom f] : lift_hom₂ (λ_ f).hom := +{ lift := (λ_ (lift_hom.lift f)).hom, } +instance lift_hom₂_left_unitor_inv (f : a ⟶ b) [lift_hom f] : lift_hom₂ (λ_ f).inv := +{ lift := (λ_ (lift_hom.lift f)).inv, } +instance lift_hom₂_right_unitor_hom (f : a ⟶ b) [lift_hom f] : lift_hom₂ (ρ_ f).hom := +{ lift := (ρ_ (lift_hom.lift f)).hom, } +instance lift_hom₂_right_unitor_inv (f : a ⟶ b) [lift_hom f] : lift_hom₂ (ρ_ f).inv := +{ lift := (ρ_ (lift_hom.lift f)).inv, } +instance lift_hom₂_associator_hom (f : a ⟶ b) (g : b ⟶ c) (h : c ⟶ d) + [lift_hom f] [lift_hom g] [lift_hom h] : + lift_hom₂ (α_ f g h).hom := +{ lift := (α_ (lift_hom.lift f) (lift_hom.lift g) (lift_hom.lift h)).hom, } +instance lift_hom₂_associator_inv (f : a ⟶ b) (g : b ⟶ c) (h : c ⟶ d) + [lift_hom f] [lift_hom g] [lift_hom h] : + lift_hom₂ (α_ f g h).inv := +{ lift := (α_ (lift_hom.lift f) (lift_hom.lift g) (lift_hom.lift h)).inv, } +instance lift_hom₂_comp {f g h : a ⟶ b} [lift_hom f] [lift_hom g] [lift_hom h] + (η : f ⟶ g) (θ : g ⟶ h) [lift_hom₂ η] [lift_hom₂ θ] : lift_hom₂ (η ≫ θ) := +{ lift := lift_hom₂.lift η ≫ lift_hom₂.lift θ } +instance lift_hom₂_whisker_left (f : a ⟶ b) [lift_hom f] {g h : b ⟶ c} (η : g ⟶ h) + [lift_hom g] [lift_hom h] [lift_hom₂ η] : lift_hom₂ (f ◁ η) := +{ lift := lift_hom.lift f ◁ lift_hom₂.lift η } +instance lift_hom₂_whisker_right {f g : a ⟶ b} (η : f ⟶ g) [lift_hom f] [lift_hom g] [lift_hom₂ η] + {h : b ⟶ c} [lift_hom h] : lift_hom₂ (η ▷ h) := +{ lift := lift_hom₂.lift η ▷ lift_hom.lift h } + +/-- +A typeclass carrying a choice of bicategorical structural isomorphism between two objects. +Used by the `⊗≫` bicategorical composition operator, and the `coherence` tactic. +-/ +class bicategorical_coherence (f g : a ⟶ b) [lift_hom f] [lift_hom g] := +(hom [] : f ⟶ g) +[is_iso : is_iso hom . tactic.apply_instance] + +attribute [instance] bicategorical_coherence.is_iso + +namespace bicategorical_coherence + +@[simps] +instance refl (f : a ⟶ b) [lift_hom f] : bicategorical_coherence f f := ⟨𝟙 _⟩ + +@[simps] +instance whisker_left + (f : a ⟶ b) (g h : b ⟶ c) [lift_hom f][lift_hom g] [lift_hom h] [bicategorical_coherence g h] : + bicategorical_coherence (f ≫ g) (f ≫ h) := +⟨f ◁ bicategorical_coherence.hom g h⟩ + +@[simps] +instance whisker_right + (f g : a ⟶ b) (h : b ⟶ c) [lift_hom f] [lift_hom g] [lift_hom h] [bicategorical_coherence f g] : + bicategorical_coherence (f ≫ h) (g ≫ h) := +⟨bicategorical_coherence.hom f g ▷ h⟩ + +@[simps] +instance tensor_right (f : a ⟶ b) (g : b ⟶ b) [lift_hom f] [lift_hom g] + [bicategorical_coherence (𝟙 b) g] : + bicategorical_coherence f (f ≫ g) := +⟨(ρ_ f).inv ≫ (f ◁ bicategorical_coherence.hom (𝟙 b) g)⟩ + +@[simps] +instance tensor_right' (f : a ⟶ b) (g : b ⟶ b) [lift_hom f] [lift_hom g] + [bicategorical_coherence g (𝟙 b)] : + bicategorical_coherence (f ≫ g) f := +⟨(f ◁ bicategorical_coherence.hom g (𝟙 b)) ≫ (ρ_ f).hom⟩ + +@[simps] +instance left (f g : a ⟶ b) [lift_hom f] [lift_hom g] [bicategorical_coherence f g] : + bicategorical_coherence (𝟙 a ≫ f) g := +⟨(λ_ f).hom ≫ bicategorical_coherence.hom f g⟩ + +@[simps] +instance left' (f g : a ⟶ b) [lift_hom f] [lift_hom g] [bicategorical_coherence f g] : + bicategorical_coherence f (𝟙 a ≫ g) := +⟨bicategorical_coherence.hom f g ≫ (λ_ g).inv⟩ + +@[simps] +instance right (f g : a ⟶ b) [lift_hom f] [lift_hom g] [bicategorical_coherence f g] : + bicategorical_coherence (f ≫ 𝟙 b) g := +⟨(ρ_ f).hom ≫ bicategorical_coherence.hom f g⟩ + +@[simps] +instance right' (f g : a ⟶ b) [lift_hom f] [lift_hom g] [bicategorical_coherence f g] : + bicategorical_coherence f (g ≫ 𝟙 b) := +⟨bicategorical_coherence.hom f g ≫ (ρ_ g).inv⟩ + +@[simps] +instance assoc (f : a ⟶ b) (g : b ⟶ c) (h : c ⟶ d) (i : a ⟶ d) + [lift_hom f] [lift_hom g] [lift_hom h] [lift_hom i] [bicategorical_coherence (f ≫ (g ≫ h)) i] : + bicategorical_coherence ((f ≫ g) ≫ h) i := +⟨(α_ f g h).hom ≫ bicategorical_coherence.hom (f ≫ (g ≫ h)) i⟩ + +@[simps] +instance assoc' (f : a ⟶ b) (g : b ⟶ c) (h : c ⟶ d) (i : a ⟶ d) + [lift_hom f] [lift_hom g] [lift_hom h] [lift_hom i] [bicategorical_coherence i (f ≫ (g ≫ h))] : + bicategorical_coherence i ((f ≫ g) ≫ h) := +⟨bicategorical_coherence.hom i (f ≫ (g ≫ h)) ≫ (α_ f g h).inv⟩ + +end bicategorical_coherence + +/-- Construct an isomorphism between two objects in a bicategorical category +out of unitors and associators. -/ +def bicategorical_iso (f g : a ⟶ b) [lift_hom f] [lift_hom g] [bicategorical_coherence f g] : + f ≅ g := +as_iso (bicategorical_coherence.hom f g) + +/-- Compose two morphisms in a bicategorical category, +inserting unitors and associators between as necessary. -/ +def bicategorical_comp {f g h i : a ⟶ b} [lift_hom g] [lift_hom h] + [bicategorical_coherence g h] (η : f ⟶ g) (θ : h ⟶ i) : f ⟶ i := +η ≫ bicategorical_coherence.hom g h ≫ θ + +localized "infixr ` ⊗≫ `:80 := category_theory.bicategory.bicategorical_comp" + in bicategory -- type as \ot \gg + +/-- Compose two isomorphisms in a bicategorical category, +inserting unitors and associators between as necessary. -/ +def bicategorical_iso_comp {f g h i : a ⟶ b} [lift_hom g] [lift_hom h] + [bicategorical_coherence g h] (η : f ≅ g) (θ : h ≅ i) : f ≅ i := +η ≪≫ as_iso (bicategorical_coherence.hom g h) ≪≫ θ + +localized "infixr ` ≪⊗≫ `:80 := category_theory.bicategory.bicategorical_iso_comp" + in bicategory -- type as \ot \gg + +example {f' : a ⟶ d} {f : a ⟶ b} {g : b ⟶ c} {h : c ⟶ d} {h' : a ⟶ d} + (η : f' ⟶ f ≫ (g ≫ h)) (θ : (f ≫ g) ≫ h ⟶ h') : f' ⟶ h' := η ⊗≫ θ + +-- To automatically insert unitors/associators at the beginning or end, +-- you can use `η ⊗≫ 𝟙 _` +example {f' : a ⟶ d } {f : a ⟶ b} {g : b ⟶ c} {h : c ⟶ d} (η : f' ⟶ (f ≫ g) ≫ h) : + f' ⟶ f ≫ (g ≫ h) := η ⊗≫ 𝟙 _ + +@[simp] lemma bicategorical_comp_refl {f g h : a ⟶ b} (η : f ⟶ g) (θ : g ⟶ h) : + η ⊗≫ θ = η ≫ θ := +by { dsimp [bicategorical_comp], simp, } + +end category_theory.bicategory + +open category_theory.bicategory + +namespace tactic + +setup_tactic_parser + +/-- Coherence tactic for bicategories. -/ +meta def bicategorical_coherence : tactic unit := +focus1 $ +do + o ← get_options, set_options $ o.set_nat `class.instance_max_depth 128, + try `[dsimp], + `(%%lhs = %%rhs) ← target, + to_expr ``((free_bicategory.lift (prefunctor.id _)).map₂ (lift_hom₂.lift %%lhs) = + (free_bicategory.lift (prefunctor.id _)).map₂ (lift_hom₂.lift %%rhs)) + >>= tactic.change, + congr + +namespace bicategory + +/-- Simp lemmas for rewriting a 2-morphism into a normal form. -/ +meta def whisker_simps : tactic unit := +`[simp only [ + category_theory.category.assoc, + category_theory.bicategory.comp_whisker_left, + category_theory.bicategory.id_whisker_left, + category_theory.bicategory.whisker_right_comp, + category_theory.bicategory.whisker_right_id, + category_theory.bicategory.whisker_left_comp, + category_theory.bicategory.whisker_left_id, + category_theory.bicategory.comp_whisker_right, + category_theory.bicategory.id_whisker_right, + category_theory.bicategory.whisker_assoc]] + +namespace coherence + +/-- +Auxiliary simp lemma for the `coherence` tactic: +this move brackets to the left in order to expose a maximal prefix +built out of unitors and associators. +-/ +-- We have unused typeclass arguments here. +-- They are intentional, to ensure that `simp only [assoc_lift_hom₂]` only left associates +-- bicategorical structural morphisms. +@[nolint unused_arguments] +lemma assoc_lift_hom₂ {f g h i : a ⟶ b} [lift_hom f] [lift_hom g] [lift_hom h] + (η : f ⟶ g) (θ : g ⟶ h) (ι : h ⟶ i) [lift_hom₂ η] [lift_hom₂ θ] : + η ≫ (θ ≫ ι) = (η ≫ θ) ≫ ι := +(category.assoc _ _ _).symm + +end coherence + +end bicategory + +end tactic diff --git a/src/category_theory/monoidal/coherence.lean b/src/category_theory/monoidal/coherence.lean index bc1e0775c9553..27b7e7bc6b669 100644 --- a/src/category_theory/monoidal/coherence.lean +++ b/src/category_theory/monoidal/coherence.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison, Yuma Mizuno, Oleksandr Manzyuk -/ import category_theory.monoidal.free.coherence +import category_theory.bicategory.coherence_tactic /-! # A `coherence` tactic for monoidal categories, and `⊗≫` (composition up to associators) @@ -217,8 +218,7 @@ Users will typicall just use the `coherence` tactic, which can also cope with id `a ≫ f ≫ b ≫ g ≫ c = a' ≫ f ≫ b' ≫ g ≫ c'` where `a = a'`, `b = b'`, and `c = c'` can be proved using `pure_coherence` -/ --- TODO: provide the `bicategory_coherence` tactic, and add that here. -meta def pure_coherence : tactic unit := monoidal_coherence +meta def pure_coherence : tactic unit := monoidal_coherence <|> bicategorical_coherence example (X₁ X₂ : C) : ((λ_ (𝟙_ C)).inv ⊗ 𝟙 (X₁ ⊗ X₂)) ≫ (α_ (𝟙_ C) (𝟙_ C) (X₁ ⊗ X₂)).hom ≫ @@ -254,7 +254,7 @@ do o ← get_options, set_options $ o.set_nat `class.instance_max_depth 128, try `[simp only [monoidal_comp, category_theory.category.assoc]] >> `[apply (cancel_epi (𝟙 _)).1; try { apply_instance }] >> - try `[simp only [tactic.coherence.assoc_lift_hom]] + try `[simp only [tactic.coherence.assoc_lift_hom, tactic.bicategory.coherence.assoc_lift_hom₂]] example {W X Y Z : C} (f : Y ⟶ Z) (g) (w : false) : (λ_ _).hom ≫ f = g := begin @@ -272,21 +272,8 @@ end coherence open coherence -/-- -Use the coherence theorem for monoidal categories to solve equations in a monoidal equation, -where the two sides only differ by replacing strings of monoidal structural morphisms -(that is, associators, unitors, and identities) -with different strings of structural morphisms with the same source and target. - -That is, `coherence` can handle goals of the form -`a ≫ f ≫ b ≫ g ≫ c = a' ≫ f ≫ b' ≫ g ≫ c'` -where `a = a'`, `b = b'`, and `c = c'` can be proved using `pure_coherence`. - -(If you have very large equations on which `coherence` is unexpectedly failing, -you may need to increase the typeclass search depth, -using e.g. `set_option class.instance_max_depth 500`.) --/ -meta def coherence : tactic unit := +/-- The main part of `coherence` tactic. -/ +meta def coherence_loop : tactic unit := do -- To prove an equality `f = g` in a monoidal category, -- first try the `pure_coherence` tactic on the entire equation: @@ -313,7 +300,29 @@ do -- with identical first terms, reflexivity <|> fail "`coherence` tactic failed, non-structural morphisms don't match", -- and whose second terms can be identified by recursively called `coherence`. - coherence) + coherence_loop) + +/-- +Use the coherence theorem for monoidal categories to solve equations in a monoidal equation, +where the two sides only differ by replacing strings of monoidal structural morphisms +(that is, associators, unitors, and identities) +with different strings of structural morphisms with the same source and target. + +That is, `coherence` can handle goals of the form +`a ≫ f ≫ b ≫ g ≫ c = a' ≫ f ≫ b' ≫ g ≫ c'` +where `a = a'`, `b = b'`, and `c = c'` can be proved using `pure_coherence`. + +(If you have very large equations on which `coherence` is unexpectedly failing, +you may need to increase the typeclass search depth, +using e.g. `set_option class.instance_max_depth 500`.) +-/ +meta def coherence : tactic unit := +do + try `[simp only [bicategorical_comp]], + try `[simp only [monoidal_comp]], + -- TODO: put similar normalization simp lemmas for monoidal categories + try bicategory.whisker_simps, + coherence_loop run_cmd add_interactive [`pure_coherence, `coherence] diff --git a/src/category_theory/monoidal/internal/types.lean b/src/category_theory/monoidal/internal/types.lean index 097b6a5d58712..0830caa718827 100644 --- a/src/category_theory/monoidal/internal/types.lean +++ b/src/category_theory/monoidal/internal/types.lean @@ -47,6 +47,8 @@ def inverse : Mon.{u} ⥤ Mon_ (Type u) := { X := A, one := λ _, 1, mul := λ p, p.1 * p.2, + one_mul' := by { ext ⟨_, _⟩, dsimp, simp, }, + mul_one' := by { ext ⟨_, _⟩, dsimp, simp, }, mul_assoc' := by { ext ⟨⟨x, y⟩, z⟩, simp [mul_assoc], }, }, map := λ A B f, { hom := f, }, } diff --git a/test/coherence.lean b/test/coherence.lean index 263f4ee9414c7..2d34bec837e6b 100644 --- a/test/coherence.lean +++ b/test/coherence.lean @@ -4,40 +4,44 @@ open category_theory universes w v u --- section bicategory --- open_locale bicategory --- variables {B : Type u} [bicategory.{w v} B] {a b c d e : B} +section bicategory +open_locale bicategory +variables {B : Type u} [bicategory.{w v} B] {a b c d e : B} --- example : (λ_ (𝟙 a)).hom = (ρ_ (𝟙 a)).hom := by coherence --- example : (λ_ (𝟙 a)).inv = (ρ_ (𝟙 a)).inv := by coherence --- example (f : a ⟶ b) (g : b ⟶ c) (h : c ⟶ d) : --- (α_ f g h).inv ≫ (α_ f g h).hom = 𝟙 (f ≫ g ≫ h) := --- by coherence --- example (f : a ⟶ b) (g : b ⟶ c) (h : c ⟶ d) (i : d ⟶ e) : --- (f ◁ (α_ g h i).hom) ≫ (α_ f g (h ≫ i)).inv ≫ (α_ (f ≫ g) h i).inv = --- (α_ f (g ≫ h) i).inv ≫ ((α_ f g h).inv ▷ i) := --- by coherence --- example (f : a ⟶ b) (g : b ⟶ c) : --- (f ◁ (λ_ g).inv) ≫ (α_ f (𝟙 b) g).inv = (ρ_ f).inv ▷ g := --- by coherence --- example (f₁ : a ⟶ b) (g₁ : b ⟶ a) (f₂ : b ⟶ c) (g₂ : c ⟶ b) : --- (α_ (𝟙 a) (𝟙 a) (f₁ ≫ f₂)).hom ≫ --- (𝟙 a ◁ (α_ (𝟙 a) f₁ f₂).inv) ≫ --- (𝟙 a ◁ (λ_ f₁).hom ≫ (ρ_ f₁).inv ▷ f₂) ≫ --- (𝟙 a ◁ (α_ f₁ (𝟙 b) f₂).hom) ≫ --- (α_ (𝟙 a) f₁ (𝟙 b ≫ f₂)).inv ≫ --- ((λ_ f₁).hom ≫ (ρ_ f₁).inv ▷ 𝟙 b ≫ f₂) ≫ --- (α_ f₁ (𝟙 b) (𝟙 b ≫ f₂)).hom ≫ --- (f₁ ◁ 𝟙 b ◁ (λ_ f₂).hom ≫ (ρ_ f₂).inv) ≫ --- (f₁ ◁ (α_ (𝟙 b) f₂ (𝟙 c)).inv) ≫ --- (f₁ ◁ (λ_ f₂).hom ≫ (ρ_ f₂).inv ▷ 𝟙 c) ≫ --- (f₁ ◁ (α_ f₂ (𝟙 c) (𝟙 c)).hom) ≫ --- (α_ f₁ f₂ (𝟙 c ≫ 𝟙 c)).inv = --- (((λ_ (𝟙 a)).hom ▷ f₁ ≫ f₂) ≫ (λ_ (f₁ ≫ f₂)).hom ≫ (ρ_ (f₁ ≫ f₂)).inv) ≫ --- (f₁ ≫ f₂ ◁ (λ_ (𝟙 c)).inv) := --- by coherence +example : (λ_ (𝟙 a)).hom = (ρ_ (𝟙 a)).hom := by coherence +example : (λ_ (𝟙 a)).inv = (ρ_ (𝟙 a)).inv := by coherence +example (f : a ⟶ b) (g : b ⟶ c) (h : c ⟶ d) : + (α_ f g h).inv ≫ (α_ f g h).hom = 𝟙 (f ≫ g ≫ h) := +by coherence +example (f : a ⟶ b) (g : b ⟶ c) (h : c ⟶ d) (i : d ⟶ e) : + f ◁ (α_ g h i).hom ≫ (α_ f g (h ≫ i)).inv ≫ (α_ (f ≫ g) h i).inv = + (α_ f (g ≫ h) i).inv ≫ (α_ f g h).inv ▷ i := +by coherence +example (f : a ⟶ b) (g : b ⟶ c) : + f ◁ (λ_ g).inv ≫ (α_ f (𝟙 b) g).inv = (ρ_ f).inv ▷ g := +by coherence +example (f g : a ⟶ a) (η : 𝟙 a ⟶ f) (θ : f ⟶ g) (w : false) : + (λ_ (𝟙 a)).hom ≫ η ≫ 𝟙 f ≫ θ = (ρ_ (𝟙 a)).hom ≫ η ≫ θ := +by coherence + +example (f₁ : a ⟶ b) (g₁ : b ⟶ a) (f₂ : b ⟶ c) (g₂ : c ⟶ b) : + (α_ (𝟙 a) (𝟙 a) (f₁ ≫ f₂)).hom ≫ + 𝟙 a ◁ (α_ (𝟙 a) f₁ f₂).inv ≫ + 𝟙 a ◁ ((λ_ f₁).hom ≫ (ρ_ f₁).inv) ▷ f₂ ≫ + 𝟙 a ◁ (α_ f₁ (𝟙 b) f₂).hom ≫ + (α_ (𝟙 a) f₁ (𝟙 b ≫ f₂)).inv ≫ + ((λ_ f₁).hom ≫ (ρ_ f₁).inv) ▷ (𝟙 b ≫ f₂) ≫ + (α_ f₁ (𝟙 b) (𝟙 b ≫ f₂)).hom ≫ + f₁ ◁ 𝟙 b ◁ ((λ_ f₂).hom ≫ (ρ_ f₂).inv) ≫ + f₁ ◁ (α_ (𝟙 b) f₂ (𝟙 c)).inv ≫ + f₁ ◁ ((λ_ f₂).hom ≫ (ρ_ f₂).inv) ▷ 𝟙 c ≫ + (f₁ ◁ (α_ f₂ (𝟙 c) (𝟙 c)).hom) ≫ + (α_ f₁ f₂ (𝟙 c ≫ 𝟙 c)).inv = + ((λ_ (𝟙 a)).hom ▷ (f₁ ≫ f₂) ≫ (λ_ (f₁ ≫ f₂)).hom ≫ (ρ_ (f₁ ≫ f₂)).inv) ≫ + (f₁ ≫ f₂) ◁ (λ_ (𝟙 c)).inv := +by coherence --- end bicategory +end bicategory section monoidal variables {C : Type u} [category.{v} C] [monoidal_category C] From 171825a926fa0bb43dc23efcd97442322304d0db Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 4 May 2022 06:42:35 +0000 Subject: [PATCH 066/153] chore(algebra/order/floor): missing simp lemmas on floor of nat and int (#13904) --- src/algebra/order/floor.lean | 10 ++++++++++ 1 file changed, 10 insertions(+) diff --git a/src/algebra/order/floor.lean b/src/algebra/order/floor.lean index d3e49d8fe92a4..c93af3024e19f 100644 --- a/src/algebra/order/floor.lean +++ b/src/algebra/order/floor.lean @@ -79,6 +79,9 @@ def floor : α → ℕ := floor_semiring.floor /-- `⌈a⌉₊` is the least natural `n` such that `a ≤ n` -/ def ceil : α → ℕ := floor_semiring.ceil +@[simp] lemma floor_nat : (nat.floor : ℕ → ℕ) = id := rfl +@[simp] lemma ceil_nat : (nat.ceil : ℕ → ℕ) = id := rfl + notation `⌊` a `⌋₊` := nat.floor a notation `⌈` a `⌉₊` := nat.ceil a @@ -383,6 +386,10 @@ def ceil : α → ℤ := floor_ring.ceil /-- `int.fract a`, the fractional part of `a`, is `a` minus its floor. -/ def fract (a : α) : α := a - floor a +@[simp] lemma floor_int : (int.floor : ℤ → ℤ) = id := rfl +@[simp] lemma ceil_int : (int.ceil : ℤ → ℤ) = id := rfl +@[simp] lemma fract_int : (int.fract : ℤ → ℤ) = 0 := funext $ λ x, by simp [fract] + notation `⌊` a `⌋` := int.floor a notation `⌈` a `⌉` := int.ceil a -- Mathematical notation for `fract a` is usually `{a}`. Let's not even go there. @@ -689,6 +696,9 @@ lemma int.floor_to_nat (a : α) : ⌊a⌋.to_nat = ⌊a⌋₊ := rfl lemma int.ceil_to_nat (a : α) : ⌈a⌉.to_nat = ⌈a⌉₊ := rfl +@[simp] lemma nat.floor_int : (nat.floor : ℤ → ℕ) = int.to_nat := rfl +@[simp] lemma nat.ceil_int : (nat.ceil : ℤ → ℕ) = int.to_nat := rfl + variables {a : α} lemma nat.cast_floor_eq_int_floor (ha : 0 ≤ a) : (⌊a⌋₊ : ℤ) = ⌊a⌋ := From 6f3426cf4f862fbd800ca9ab8c494ceb155974e0 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Wed, 4 May 2022 06:42:36 +0000 Subject: [PATCH 067/153] chore(number_theory/legendre_symbol/quadratic_char): golf some proofs (#13926) --- .../legendre_symbol/quadratic_char.lean | 38 ++++++++----------- 1 file changed, 15 insertions(+), 23 deletions(-) diff --git a/src/number_theory/legendre_symbol/quadratic_char.lean b/src/number_theory/legendre_symbol/quadratic_char.lean index 6ccd8acff84f8..e405fada8884b 100644 --- a/src/number_theory/legendre_symbol/quadratic_char.lean +++ b/src/number_theory/legendre_symbol/quadratic_char.lean @@ -46,7 +46,7 @@ lemma odd_card_of_char_ne_two (hF : ring_char F ≠ 2) : fintype.card F % 2 = 1 begin rcases finite_field.card F (ring_char F) with ⟨ n, hp, h ⟩, have h₁ : odd ((ring_char F) ^ (n : ℕ)) := - odd.pow ((or_iff_right hF).mp (nat.prime.eq_two_or_odd' hp)), + odd.pow ((or_iff_right hF).mp (nat.prime.eq_two_or_odd' hp)), rwa [← h, nat.odd_iff] at h₁, end @@ -62,10 +62,10 @@ end lemma neg_ne_self_of_char_ne_two (hF : ring_char F ≠ 2) {a : F} (ha : a ≠ 0) : a ≠ -a := begin intro hf, - rw [eq_neg_iff_add_eq_zero, (by ring : a + a = 2 * a), mul_eq_zero] at hf, - have h := mt eq_neg_iff_add_eq_zero.mpr (neg_one_ne_one_of_char_ne_two hF).symm, - norm_num at h, - exact or.dcases_on hf (λ (hf : 2 = 0), h hf) (λ (hf : a = 0), ha hf), + apply (neg_one_ne_one_of_char_ne_two hF).symm, + rw [eq_neg_iff_add_eq_zero, ←two_mul, mul_one], + rw [eq_neg_iff_add_eq_zero, ←two_mul, mul_eq_zero] at hf, + exact hf.resolve_right ha, end /-- If `F` has odd characteristic, then for nonzero `a : F`, we have that `a ^ (#F / 2) = ±1`. -/ @@ -277,11 +277,9 @@ lemma quadratic_char_dichotomy {a : F} (ha : a ≠ 0) : lemma quadratic_char_eq_neg_one_iff_not_one {a : F} (ha : a ≠ 0) : quadratic_char F a = -1 ↔ ¬ quadratic_char F a = 1 := begin - split, - { intro h, - rw h, - norm_num, }, - { exact λ h₂, (or_iff_right h₂).mp (quadratic_char_dichotomy ha), } + refine ⟨λ h, _, λ h₂, (or_iff_right h₂).mp (quadratic_char_dichotomy ha)⟩, + rw h, + norm_num, end /-- For `a : F`, `quadratic_char F a = -1 ↔ ¬ is_square a`. -/ @@ -290,14 +288,12 @@ lemma quadratic_char_neg_one_iff_not_is_square {a : F} : begin by_cases ha : a = 0, { simp only [ha, is_square_zero, quadratic_char_zero, zero_eq_neg, one_ne_zero, not_true], }, - { rw quadratic_char_eq_neg_one_iff_not_one ha, - exact not_iff_not.mpr (quadratic_char_one_iff_is_square ha), }, + { rw [quadratic_char_eq_neg_one_iff_not_one ha, quadratic_char_one_iff_is_square ha] }, end /-- If `F` has odd characteristic, then `quadratic_char F` takes the value `-1`. -/ lemma quadratic_char_exists_neg_one (hF : ring_char F ≠ 2) : ∃ a, quadratic_char F a = -1 := -Exists.dcases_on (finite_field.exists_nonsquare hF) - (λ (b : F) (h₁ : ¬is_square b), ⟨b, (quadratic_char_neg_one_iff_not_is_square.mpr h₁)⟩) +(finite_field.exists_nonsquare hF).imp (λ b h₁, quadratic_char_neg_one_iff_not_is_square.mpr h₁) /-- The number of solutions to `x^2 = a` is determined by the quadratic character. -/ lemma quadratic_char_card_sqrts (hF : ring_char F ≠ 2) (a : F) : @@ -312,20 +308,18 @@ begin by_cases h : is_square a, { rw (quadratic_char_one_iff_is_square h₀).mpr h, rcases h with ⟨b, h⟩, - have hb₀ : b ≠ 0 := by { intro hb, rw [hb, mul_zero] at h, exact h₀ h, }, + rw [h, mul_self_eq_zero] at h₀, have h₁ : s = [b, -b].to_finset := by { ext x, simp only [finset.mem_filter, finset.mem_univ, true_and, list.to_finset_cons, list.to_finset_nil, insert_emptyc_eq, finset.mem_insert, finset.mem_singleton], rw ← pow_two at h, - rw hs, - simp only [set.mem_to_finset, set.mem_set_of_eq], - rw h, + simp only [hs, set.mem_to_finset, set.mem_set_of_eq, h], split, { exact eq_or_eq_neg_of_sq_eq_sq _ _, }, { rintro (h₂ | h₂); rw h₂, simp only [neg_sq], }, }, - simp only [h₁, finset.card_doubleton (finite_field.neg_ne_self_of_char_ne_two hF hb₀), + simp only [h₁, finset.card_doubleton (finite_field.neg_ne_self_of_char_ne_two hF h₀), list.to_finset_cons, list.to_finset_nil, insert_emptyc_eq, int.coe_nat_succ, int.coe_nat_zero, zero_add], }, { rw quadratic_char_neg_one_iff_not_is_square.mpr h, @@ -334,8 +328,7 @@ begin ext x, simp only [iff_false, finset.mem_filter, finset.mem_univ, true_and, finset.not_mem_empty], rw is_square_iff_exists_sq at h, - push_neg at h, - exact (h x).symm, }, }, + exact λ h', h ⟨_, h'.symm⟩, }, }, end open_locale big_operators @@ -348,9 +341,8 @@ begin { intro hf, rw [hf, quadratic_char_zero, zero_eq_neg] at hb, exact one_ne_zero hb, }, - let mul_b : F → F := λ x, b * x, have h₁ : ∑ (a : F), quadratic_char F (b * a) = ∑ (a : F), quadratic_char F a := - by refine fintype.sum_bijective _ (mul_left_bijective₀ b h₀) _ _ (λ x, rfl), + fintype.sum_bijective _ (mul_left_bijective₀ b h₀) _ _ (λ x, rfl), simp only [quadratic_char_mul] at h₁, rw [← finset.mul_sum, hb, neg_mul, one_mul] at h₁, exact eq_zero_of_neg_eq h₁, From 58de2a0a0b901e0b5263635622177589773b1b50 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 4 May 2022 06:42:38 +0000 Subject: [PATCH 068/153] chore(analysis): use nnnorm notation everywhere (#13930) This was done with a series of ad-hoc regular expressions, then cleaned up by hand. --- src/analysis/analytic/composition.lean | 32 +++++++++---------- src/analysis/analytic/radius_liminf.lean | 4 +-- src/analysis/calculus/inverse.lean | 6 ++-- src/analysis/normed/group/hom.lean | 2 +- src/analysis/normed_space/banach.lean | 2 +- src/analysis/normed_space/enorm.lean | 14 ++++---- .../normed_space/finite_dimension.lean | 4 +-- .../normed_space/indicator_function.lean | 2 +- .../normed_space/linear_isometry.lean | 4 +-- src/analysis/normed_space/multilinear.lean | 4 +-- src/analysis/normed_space/operator_norm.lean | 8 ++--- .../constructions/borel_space.lean | 10 +++--- src/measure_theory/constructions/prod.lean | 8 ++--- .../function/conditional_expectation.lean | 2 +- src/measure_theory/function/l1_space.lean | 32 +++++++++---------- src/measure_theory/function/l2_space.lean | 2 +- src/measure_theory/function/lp_space.lean | 2 +- .../function/simple_func_dense_lp.lean | 4 +-- .../function/strongly_measurable.lean | 8 ++--- src/measure_theory/integral/set_to_l1.lean | 2 +- 20 files changed, 76 insertions(+), 76 deletions(-) diff --git a/src/analysis/analytic/composition.lean b/src/analysis/analytic/composition.lean index 603a0ba30de12..73a7fce074e95 100644 --- a/src/analysis/analytic/composition.lean +++ b/src/analysis/analytic/composition.lean @@ -256,7 +256,7 @@ multilinear_map.mk_continuous_norm_le _ lemma comp_along_composition_nnnorm {n : ℕ} (q : formal_multilinear_series 𝕜 F G) (p : formal_multilinear_series 𝕜 E F) (c : composition n) : - nnnorm (q.comp_along_composition p c) ≤ nnnorm (q c.length) * ∏ i, nnnorm (p (c.blocks_fun i)) := + ∥q.comp_along_composition p c∥₊ ≤ ∥q c.length∥₊ * ∏ i, ∥p (c.blocks_fun i)∥₊ := by { rw ← nnreal.coe_le_coe, push_cast, exact q.comp_along_composition_norm p c } /-- Formal composition of two formal multilinear series. The `n`-th coefficient in the composition @@ -432,7 +432,7 @@ theorem comp_summable_nnreal (q : formal_multilinear_series 𝕜 F G) (p : formal_multilinear_series 𝕜 E F) (hq : 0 < q.radius) (hp : 0 < p.radius) : ∃ r > (0 : ℝ≥0), - summable (λ i : Σ n, composition n, nnnorm (q.comp_along_composition p i.2) * r ^ i.1) := + summable (λ i : Σ n, composition n, ∥q.comp_along_composition p i.2∥₊ * r ^ i.1) := begin /- This follows from the fact that the growth rate of `∥qₙ∥` and `∥pₙ∥` is at most geometric, giving a geometric bound on each `∥q.comp_along_composition p op∥`, together with the @@ -440,9 +440,9 @@ begin rcases ennreal.lt_iff_exists_nnreal_btwn.1 (lt_min ennreal.zero_lt_one hq) with ⟨rq, rq_pos, hrq⟩, rcases ennreal.lt_iff_exists_nnreal_btwn.1 (lt_min ennreal.zero_lt_one hp) with ⟨rp, rp_pos, hrp⟩, simp only [lt_min_iff, ennreal.coe_lt_one_iff, ennreal.coe_pos] at hrp hrq rp_pos rq_pos, - obtain ⟨Cq, hCq0, hCq⟩ : ∃ Cq > 0, ∀ n, nnnorm (q n) * rq^n ≤ Cq := + obtain ⟨Cq, hCq0, hCq⟩ : ∃ Cq > 0, ∀ n, ∥q n∥₊ * rq^n ≤ Cq := q.nnnorm_mul_pow_le_of_lt_radius hrq.2, - obtain ⟨Cp, hCp1, hCp⟩ : ∃ Cp ≥ 1, ∀ n, nnnorm (p n) * rp^n ≤ Cp, + obtain ⟨Cp, hCp1, hCp⟩ : ∃ Cp ≥ 1, ∀ n, ∥p n∥₊ * rp^n ≤ Cp, { rcases p.nnnorm_mul_pow_le_of_lt_radius hrp.2 with ⟨Cp, -, hCp⟩, exact ⟨max Cp 1, le_max_right _ _, λ n, (hCp n).trans (le_max_left _ _)⟩ }, let r0 : ℝ≥0 := (4 * Cp)⁻¹, @@ -450,23 +450,23 @@ begin set r : ℝ≥0 := rp * rq * r0, have r_pos : 0 < r := mul_pos (mul_pos rp_pos rq_pos) r0_pos, have I : ∀ (i : Σ (n : ℕ), composition n), - nnnorm (q.comp_along_composition p i.2) * r ^ i.1 ≤ Cq / 4 ^ i.1, + ∥q.comp_along_composition p i.2∥₊ * r ^ i.1 ≤ Cq / 4 ^ i.1, { rintros ⟨n, c⟩, have A, - calc nnnorm (q c.length) * rq ^ n ≤ nnnorm (q c.length)* rq ^ c.length : + calc ∥q c.length∥₊ * rq ^ n ≤ ∥q c.length∥₊* rq ^ c.length : mul_le_mul' le_rfl (pow_le_pow_of_le_one rq.2 hrq.1.le c.length_le) ... ≤ Cq : hCq _, have B, - calc ((∏ i, nnnorm (p (c.blocks_fun i))) * rp ^ n) - = ∏ i, nnnorm (p (c.blocks_fun i)) * rp ^ c.blocks_fun i : + calc ((∏ i, ∥p (c.blocks_fun i)∥₊) * rp ^ n) + = ∏ i, ∥p (c.blocks_fun i)∥₊ * rp ^ c.blocks_fun i : by simp only [finset.prod_mul_distrib, finset.prod_pow_eq_pow_sum, c.sum_blocks_fun] ... ≤ ∏ i : fin c.length, Cp : finset.prod_le_prod' (λ i _, hCp _) ... = Cp ^ c.length : by simp ... ≤ Cp ^ n : pow_le_pow hCp1 c.length_le, - calc nnnorm (q.comp_along_composition p c) * r ^ n - ≤ (nnnorm (q c.length) * ∏ i, nnnorm (p (c.blocks_fun i))) * r ^ n : + calc ∥q.comp_along_composition p c∥₊ * r ^ n + ≤ (∥q c.length∥₊ * ∏ i, ∥p (c.blocks_fun i)∥₊) * r ^ n : mul_le_mul' (q.comp_along_composition_nnnorm p c) le_rfl - ... = (nnnorm (q c.length) * rq ^ n) * ((∏ i, nnnorm (p (c.blocks_fun i))) * rp ^ n) * r0 ^ n : + ... = (∥q c.length∥₊ * rq ^ n) * ((∏ i, ∥p (c.blocks_fun i)∥₊) * rp ^ n) * r0 ^ n : by { simp only [r, mul_pow], ring } ... ≤ Cq * Cp ^ n * r0 ^ n : mul_le_mul' (mul_le_mul' A B) le_rfl ... = Cq / 4 ^ n : @@ -496,18 +496,18 @@ end summability over all compositions. -/ theorem le_comp_radius_of_summable (q : formal_multilinear_series 𝕜 F G) (p : formal_multilinear_series 𝕜 E F) (r : ℝ≥0) - (hr : summable (λ i : (Σ n, composition n), nnnorm (q.comp_along_composition p i.2) * r ^ i.1)) : + (hr : summable (λ i : (Σ n, composition n), ∥q.comp_along_composition p i.2∥₊ * r ^ i.1)) : (r : ℝ≥0∞) ≤ (q.comp p).radius := begin refine le_radius_of_bound_nnreal _ - (∑' i : (Σ n, composition n), nnnorm (comp_along_composition q p i.snd) * r ^ i.fst) (λ n, _), - calc nnnorm (formal_multilinear_series.comp q p n) * r ^ n ≤ - ∑' (c : composition n), nnnorm (comp_along_composition q p c) * r ^ n : + (∑' i : (Σ n, composition n), ∥comp_along_composition q p i.snd∥₊ * r ^ i.fst) (λ n, _), + calc ∥formal_multilinear_series.comp q p n∥₊ * r ^ n ≤ + ∑' (c : composition n), ∥comp_along_composition q p c∥₊ * r ^ n : begin rw [tsum_fintype, ← finset.sum_mul], exact mul_le_mul' (nnnorm_sum_le _ _) le_rfl end - ... ≤ ∑' (i : Σ (n : ℕ), composition n), nnnorm (comp_along_composition q p i.snd) * r ^ i.fst : + ... ≤ ∑' (i : Σ (n : ℕ), composition n), ∥comp_along_composition q p i.snd∥₊ * r ^ i.fst : nnreal.tsum_comp_le_tsum_of_inj hr sigma_mk_injective end diff --git a/src/analysis/analytic/radius_liminf.lean b/src/analysis/analytic/radius_liminf.lean index 14a68455567e6..ecb1d750d27da 100644 --- a/src/analysis/analytic/radius_liminf.lean +++ b/src/analysis/analytic/radius_liminf.lean @@ -27,10 +27,10 @@ variables (p : formal_multilinear_series 𝕜 E F) /-- The radius of a formal multilinear series is equal to $\liminf_{n\to\infty} \frac{1}{\sqrt[n]{∥p n∥}}$. The actual statement uses `ℝ≥0` and some coercions. -/ -lemma radius_eq_liminf : p.radius = liminf at_top (λ n, 1/((nnnorm (p n)) ^ (1 / (n : ℝ)) : ℝ≥0)) := +lemma radius_eq_liminf : p.radius = liminf at_top (λ n, 1/((∥p n∥₊) ^ (1 / (n : ℝ)) : ℝ≥0)) := begin have : ∀ (r : ℝ≥0) {n : ℕ}, 0 < n → - ((r : ℝ≥0∞) ≤ 1 / ↑(nnnorm (p n) ^ (1 / (n : ℝ))) ↔ nnnorm (p n) * r ^ n ≤ 1), + ((r : ℝ≥0∞) ≤ 1 / ↑(∥p n∥₊ ^ (1 / (n : ℝ))) ↔ ∥p n∥₊ * r ^ n ≤ 1), { intros r n hn, have : 0 < (n : ℝ) := nat.cast_pos.2 hn, conv_lhs {rw [one_div, ennreal.le_inv_iff_mul_le, ← ennreal.coe_mul, diff --git a/src/analysis/calculus/inverse.lean b/src/analysis/calculus/inverse.lean index c12d38416d1d0..4532272141c61 100644 --- a/src/analysis/calculus/inverse.lean +++ b/src/analysis/calculus/inverse.lean @@ -153,7 +153,7 @@ begin end protected lemma lipschitz (hf : approximates_linear_on f f' s c) : - lipschitz_with (nnnorm f' + c) (s.restrict f) := + lipschitz_with (∥f'∥₊ + c) (s.restrict f) := by simpa only [restrict_apply, add_sub_cancel'_right] using (f'.lipschitz.restrict s).add hf.lipschitz_sub @@ -356,7 +356,7 @@ We also assume that either `E = {0}`, or `c < ∥f'⁻¹∥⁻¹`. We use `N` as variables {f' : E ≃L[𝕜] F} {s : set E} {c : ℝ≥0} -local notation `N` := nnnorm (f'.symm : F →L[𝕜] E) +local notation `N` := ∥(f'.symm : F →L[𝕜] E)∥₊ protected lemma antilipschitz (hf : approximates_linear_on f (f' : E →L[𝕜] F) s c) (hc : subsingleton E ∨ c < N⁻¹) : @@ -572,7 +572,7 @@ variables [cs : complete_space E] {f : E → F} {f' : E ≃L[𝕜] F} {a : E} lemma approximates_deriv_on_open_nhds (hf : has_strict_fderiv_at f (f' : E →L[𝕜] F) a) : ∃ (s : set E) (hs : a ∈ s ∧ is_open s), - approximates_linear_on f (f' : E →L[𝕜] F) s ((nnnorm (f'.symm : F →L[𝕜] E))⁻¹ / 2) := + approximates_linear_on f (f' : E →L[𝕜] F) s (∥(f'.symm : F →L[𝕜] E)∥₊⁻¹ / 2) := begin refine ((nhds_basis_opens a).exists_iff _).1 _, exact (λ s t, approximates_linear_on.mono_set), diff --git a/src/analysis/normed/group/hom.lean b/src/analysis/normed/group/hom.lean index 6e79fe2426523..cefa5097fd43d 100644 --- a/src/analysis/normed/group/hom.lean +++ b/src/analysis/normed/group/hom.lean @@ -47,7 +47,7 @@ def mk_normed_group_hom (f : V →+ W) /-- Associate to a group homomorphism a bounded group homomorphism under a norm control condition. See `add_monoid_hom.mk_normed_group_hom` for a version that uses `ℝ` for the bound. -/ -def mk_normed_group_hom' (f : V →+ W) (C : ℝ≥0) (hC : ∀ x, nnnorm (f x) ≤ C * nnnorm x) : +def mk_normed_group_hom' (f : V →+ W) (C : ℝ≥0) (hC : ∀ x, ∥f x∥₊ ≤ C * ∥x∥₊) : normed_group_hom V W := { bound' := ⟨C, hC⟩ .. f} diff --git a/src/analysis/normed_space/banach.lean b/src/analysis/normed_space/banach.lean index 79753b00e498d..8a1592f87469c 100644 --- a/src/analysis/normed_space/banach.lean +++ b/src/analysis/normed_space/banach.lean @@ -52,7 +52,7 @@ end continuous_linear_map noncomputable def continuous_linear_equiv.to_nonlinear_right_inverse (f : E ≃L[𝕜] F) : continuous_linear_map.nonlinear_right_inverse (f : E →L[𝕜] F) := { to_fun := f.inv_fun, - nnnorm := nnnorm (f.symm : F →L[𝕜] E), + nnnorm := ∥(f.symm : F →L[𝕜] E)∥₊, bound' := λ y, continuous_linear_map.le_op_norm (f.symm : F →L[𝕜] E) _, right_inv' := f.apply_symm_apply } diff --git a/src/analysis/normed_space/enorm.lean b/src/analysis/normed_space/enorm.lean index fc0596712b4c8..54c986e82aed4 100644 --- a/src/analysis/normed_space/enorm.lean +++ b/src/analysis/normed_space/enorm.lean @@ -40,7 +40,7 @@ structure enorm (𝕜 : Type*) (V : Type*) [normed_field 𝕜] [add_comm_group V (to_fun : V → ℝ≥0∞) (eq_zero' : ∀ x, to_fun x = 0 → x = 0) (map_add_le' : ∀ x y : V, to_fun (x + y) ≤ to_fun x + to_fun y) -(map_smul_le' : ∀ (c : 𝕜) (x : V), to_fun (c • x) ≤ nnnorm c * to_fun x) +(map_smul_le' : ∀ (c : 𝕜) (x : V), to_fun (c • x) ≤ ∥c∥₊ * to_fun x) namespace enorm @@ -61,12 +61,12 @@ lemma ext_iff {e₁ e₂ : enorm 𝕜 V} : e₁ = e₂ ↔ ∀ x, e₁ x = e₂ @[simp, norm_cast] lemma coe_inj {e₁ e₂ : enorm 𝕜 V} : (e₁ : V → ℝ≥0∞) = e₂ ↔ e₁ = e₂ := coe_fn_injective.eq_iff -@[simp] lemma map_smul (c : 𝕜) (x : V) : e (c • x) = nnnorm c * e x := +@[simp] lemma map_smul (c : 𝕜) (x : V) : e (c • x) = ∥c∥₊ * e x := le_antisymm (e.map_smul_le' c x) $ begin by_cases hc : c = 0, { simp [hc] }, - calc (nnnorm c : ℝ≥0∞) * e x = nnnorm c * e (c⁻¹ • c • x) : by rw [inv_smul_smul₀ hc] - ... ≤ nnnorm c * (nnnorm (c⁻¹) * e (c • x)) : _ + calc (∥c∥₊ : ℝ≥0∞) * e x = ∥c∥₊ * e (c⁻¹ • c • x) : by rw [inv_smul_smul₀ hc] + ... ≤ ∥c∥₊ * (∥c⁻¹∥₊ * e (c • x)) : _ ... = e (c • x) : _, { exact ennreal.mul_le_mul le_rfl (e.map_smul_le' _ _) }, { rw [← mul_assoc, nnnorm_inv, ennreal.coe_inv, @@ -80,8 +80,8 @@ by { rw [← zero_smul 𝕜 (0:V), e.map_smul], norm_num } ⟨e.eq_zero' x, λ h, h.symm ▸ e.map_zero⟩ @[simp] lemma map_neg (x : V) : e (-x) = e x := -calc e (-x) = nnnorm (-1:𝕜) * e x : by rw [← map_smul, neg_one_smul] - ... = e x : by simp +calc e (-x) = ∥(-1 : 𝕜)∥₊ * e x : by rw [← map_smul, neg_one_smul] + ... = e x : by simp lemma map_sub_rev (x y : V) : e (x - y) = e (y - x) := by rw [← neg_sub, e.map_neg] @@ -162,7 +162,7 @@ def finite_subspace : subspace 𝕜 V := zero_mem' := by simp, add_mem' := λ x y hx hy, lt_of_le_of_lt (e.map_add_le x y) (ennreal.add_lt_top.2 ⟨hx, hy⟩), smul_mem' := λ c x (hx : _ < _), - calc e (c • x) = nnnorm c * e x : e.map_smul c x + calc e (c • x) = ∥c∥₊ * e x : e.map_smul c x ... < ⊤ : ennreal.mul_lt_top ennreal.coe_ne_top hx.ne } /-- Metric space structure on `e.finite_subspace`. We use `emetric_space.to_metric_space_of_dist` diff --git a/src/analysis/normed_space/finite_dimension.lean b/src/analysis/normed_space/finite_dimension.lean index 6e6c7a5d82953..cfdde83cd0715 100644 --- a/src/analysis/normed_space/finite_dimension.lean +++ b/src/analysis/normed_space/finite_dimension.lean @@ -425,7 +425,7 @@ begin ((continuous_apply i).tendsto _).sub tendsto_const_nhds), simp only [sub_self, norm_zero, finset.sum_const_zero] at this, refine (this.eventually (gt_mem_nhds $ inv_pos.2 K0)).mono (λ g hg, _), - replace hg : ∑ i, nnnorm (g i - f i) < K⁻¹, by { rw ← nnreal.coe_lt_coe, push_cast, exact hg }, + replace hg : ∑ i, ∥g i - f i∥₊ < K⁻¹, by { rw ← nnreal.coe_lt_coe, push_cast, exact hg }, rw linear_map.ker_eq_bot, refine (hK.add_sub_lipschitz_with (lipschitz_with.of_dist_le_mul $ λ v u, _) hg).injective, simp only [dist_eq_norm, linear_map.lsum_apply, pi.sub_apply, linear_map.sum_apply, @@ -858,7 +858,7 @@ begin set e := v.equiv_funL, have : summable (λ x, ∥e (f x)∥) := this (e.summable.2 hf), refine summable_of_norm_bounded _ (this.mul_left - ↑(nnnorm (e.symm : (fin (finrank ℝ E) → ℝ) →L[ℝ] E))) (λ i, _), + ↑(∥(e.symm : (fin (finrank ℝ E) → ℝ) →L[ℝ] E)∥₊)) (λ i, _), simpa using (e.symm : (fin (finrank ℝ E) → ℝ) →L[ℝ] E).le_op_norm (e $ f i) }, unfreezingI { clear_dependent E }, -- Now we deal with `g : α → fin N → ℝ` diff --git a/src/analysis/normed_space/indicator_function.lean b/src/analysis/normed_space/indicator_function.lean index 8dc7e9d15c95e..746e207f9b985 100644 --- a/src/analysis/normed_space/indicator_function.lean +++ b/src/analysis/normed_space/indicator_function.lean @@ -24,7 +24,7 @@ lemma norm_indicator_eq_indicator_norm : flip congr_fun a (indicator_comp_of_zero norm_zero).symm lemma nnnorm_indicator_eq_indicator_nnnorm : - nnnorm (indicator s f a) = indicator s (λa, nnnorm (f a)) a := + ∥indicator s f a∥₊ = indicator s (λa, ∥f a∥₊) a := flip congr_fun a (indicator_comp_of_zero nnnorm_zero).symm lemma norm_indicator_le_of_subset (h : s ⊆ t) (f : α → E) (a : α) : diff --git a/src/analysis/normed_space/linear_isometry.lean b/src/analysis/normed_space/linear_isometry.lean index b16391b4bf58b..38d372c854af0 100644 --- a/src/analysis/normed_space/linear_isometry.lean +++ b/src/analysis/normed_space/linear_isometry.lean @@ -102,7 +102,7 @@ f.to_linear_map.map_smul c x @[simp] lemma norm_map (x : E) : ∥f x∥ = ∥x∥ := f.norm_map' x -@[simp] lemma nnnorm_map (x : E) : nnnorm (f x) = nnnorm x := nnreal.eq $ f.norm_map x +@[simp] lemma nnnorm_map (x : E) : ∥f x∥₊ = ∥x∥₊ := nnreal.eq $ f.norm_map x protected lemma isometry : isometry f := f.to_linear_map.to_add_monoid_hom.isometry_of_norm f.norm_map @@ -493,7 +493,7 @@ omit σ₂₁ @[simp] lemma map_smul [module R E₂] {e : E ≃ₗᵢ[R] E₂} (c : R) (x : E) : e (c • x) = c • e x := e.1.map_smul c x -@[simp] lemma nnnorm_map (x : E) : nnnorm (e x) = nnnorm x := e.to_linear_isometry.nnnorm_map x +@[simp] lemma nnnorm_map (x : E) : ∥e x∥₊ = ∥x∥₊ := e.to_linear_isometry.nnnorm_map x @[simp] lemma dist_map (x y : E) : dist (e x) (e y) = dist x y := e.to_linear_isometry.dist_map x y diff --git a/src/analysis/normed_space/multilinear.lean b/src/analysis/normed_space/multilinear.lean index 81479135ec78e..366cdbac80655 100644 --- a/src/analysis/normed_space/multilinear.lean +++ b/src/analysis/normed_space/multilinear.lean @@ -408,10 +408,10 @@ by simpa only [fintype.card_fin] /-- The fundamental property of the operator norm of a continuous multilinear map: `∥f m∥` is bounded by `∥f∥` times the product of the `∥m i∥`, `nnnorm` version. -/ -theorem le_op_nnnorm : nnnorm (f m) ≤ nnnorm f * ∏ i, nnnorm (m i) := +theorem le_op_nnnorm : ∥f m∥₊ ≤ ∥f∥₊ * ∏ i, ∥m i∥₊ := nnreal.coe_le_coe.1 $ by { push_cast, exact f.le_op_norm m } -theorem le_of_op_nnnorm_le {C : ℝ≥0} (h : nnnorm f ≤ C) : nnnorm (f m) ≤ C * ∏ i, nnnorm (m i) := +theorem le_of_op_nnnorm_le {C : ℝ≥0} (h : ∥f∥₊ ≤ C) : ∥f m∥₊ ≤ C * ∏ i, ∥m i∥₊ := (f.le_op_nnnorm m).trans $ mul_le_mul' h le_rfl lemma op_norm_prod (f : continuous_multilinear_map 𝕜 E G) (g : continuous_multilinear_map 𝕜 E G') : diff --git a/src/analysis/normed_space/operator_norm.lean b/src/analysis/normed_space/operator_norm.lean index 08465fd5c898e..e0090f6bfa77f 100644 --- a/src/analysis/normed_space/operator_norm.lean +++ b/src/analysis/normed_space/operator_norm.lean @@ -1328,7 +1328,7 @@ begin exact hx.trans_lt (half_lt_self εpos) }, simpa using this }, rcases normed_field.exists_one_lt_norm 𝕜 with ⟨c, hc⟩, - refine ⟨⟨δ⁻¹, _⟩ * nnnorm c, f.to_linear_map.antilipschitz_of_bound $ λx, _⟩, + refine ⟨⟨δ⁻¹, _⟩ * ∥c∥₊, f.to_linear_map.antilipschitz_of_bound $ λx, _⟩, exact inv_nonneg.2 (le_of_lt δ_pos), by_cases hx : f x = 0, { have : f x = f 0, by { simp [hx] }, @@ -1668,7 +1668,7 @@ section variables [ring_hom_isometric σ₂₁] protected lemma antilipschitz (e : E ≃SL[σ₁₂] F) : - antilipschitz_with (nnnorm (e.symm : F →SL[σ₂₁] E)) e := + antilipschitz_with ∥(e.symm : F →SL[σ₂₁] E)∥₊ e := e.symm.lipschitz.to_right_inverse e.left_inv lemma one_le_norm_mul_norm_symm [ring_hom_isometric σ₁₂] [nontrivial E] (e : E ≃SL[σ₁₂] F) : @@ -1690,7 +1690,7 @@ lemma norm_symm_pos [ring_hom_isometric σ₁₂] [nontrivial E] (e : E ≃SL[σ pos_of_mul_pos_left (lt_of_lt_of_le zero_lt_one e.one_le_norm_mul_norm_symm) (norm_nonneg _) lemma nnnorm_symm_pos [ring_hom_isometric σ₁₂] [nontrivial E] (e : E ≃SL[σ₁₂] F) : - 0 < nnnorm (e.symm : F →SL[σ₂₁] E) := + 0 < ∥(e.symm : F →SL[σ₂₁] E)∥₊ := e.norm_symm_pos lemma subsingleton_or_norm_symm_pos [ring_hom_isometric σ₁₂] (e : E ≃SL[σ₁₂] F) : @@ -1702,7 +1702,7 @@ begin end lemma subsingleton_or_nnnorm_symm_pos [ring_hom_isometric σ₁₂] (e : E ≃SL[σ₁₂] F) : - subsingleton E ∨ 0 < (nnnorm $ (e.symm : F →SL[σ₂₁] E)) := + subsingleton E ∨ 0 < ∥(e.symm : F →SL[σ₂₁] E)∥₊ := subsingleton_or_norm_symm_pos e variable (𝕜) diff --git a/src/measure_theory/constructions/borel_space.lean b/src/measure_theory/constructions/borel_space.lean index 406e6b7dc175a..d00dd27457ff2 100644 --- a/src/measure_theory/constructions/borel_space.lean +++ b/src/measure_theory/constructions/borel_space.lean @@ -1807,26 +1807,26 @@ lemma measurable_nnnorm : measurable (nnnorm : α → ℝ≥0) := continuous_nnnorm.measurable @[measurability] -lemma measurable.nnnorm {f : β → α} (hf : measurable f) : measurable (λ a, nnnorm (f a)) := +lemma measurable.nnnorm {f : β → α} (hf : measurable f) : measurable (λ a, ∥f a∥₊) := measurable_nnnorm.comp hf @[measurability] lemma ae_measurable.nnnorm {f : β → α} {μ : measure β} (hf : ae_measurable f μ) : - ae_measurable (λ a, nnnorm (f a)) μ := + ae_measurable (λ a, ∥f a∥₊) μ := measurable_nnnorm.comp_ae_measurable hf @[measurability] -lemma measurable_ennnorm : measurable (λ x : α, (nnnorm x : ℝ≥0∞)) := +lemma measurable_ennnorm : measurable (λ x : α, (∥x∥₊ : ℝ≥0∞)) := measurable_nnnorm.coe_nnreal_ennreal @[measurability] lemma measurable.ennnorm {f : β → α} (hf : measurable f) : - measurable (λ a, (nnnorm (f a) : ℝ≥0∞)) := + measurable (λ a, (∥f a∥₊ : ℝ≥0∞)) := hf.nnnorm.coe_nnreal_ennreal @[measurability] lemma ae_measurable.ennnorm {f : β → α} {μ : measure β} (hf : ae_measurable f μ) : - ae_measurable (λ a, (nnnorm (f a) : ℝ≥0∞)) μ := + ae_measurable (λ a, (∥f a∥₊ : ℝ≥0∞)) μ := measurable_ennnorm.comp_ae_measurable hf end normed_group diff --git a/src/measure_theory/constructions/prod.lean b/src/measure_theory/constructions/prod.lean index c9ff36da8ea06..36bf4d0157c5e 100644 --- a/src/measure_theory/constructions/prod.lean +++ b/src/measure_theory/constructions/prod.lean @@ -906,14 +906,14 @@ begin rw [continuous_iff_continuous_at], intro g, refine tendsto_integral_of_L1 _ (L1.integrable_coe_fn g).integral_prod_left (eventually_of_forall $ λ h, (L1.integrable_coe_fn h).integral_prod_left) _, - simp_rw [← lintegral_fn_integral_sub (λ x, (nnnorm x : ℝ≥0∞)) (L1.integrable_coe_fn _) + simp_rw [← lintegral_fn_integral_sub (λ x, (∥x∥₊ : ℝ≥0∞)) (L1.integrable_coe_fn _) (L1.integrable_coe_fn g)], refine tendsto_of_tendsto_of_tendsto_of_le_of_le tendsto_const_nhds _ (λ i, zero_le _) _, - { exact λ i, ∫⁻ x, ∫⁻ y, nnnorm (i (x, y) - g (x, y)) ∂ν ∂μ }, + { exact λ i, ∫⁻ x, ∫⁻ y, ∥i (x, y) - g (x, y)∥₊ ∂ν ∂μ }, swap, { exact λ i, lintegral_mono (λ x, ennnorm_integral_le_lintegral_ennnorm _) }, show tendsto (λ (i : α × β →₁[μ.prod ν] E), - ∫⁻ x, ∫⁻ (y : β), nnnorm (i (x, y) - g (x, y)) ∂ν ∂μ) (𝓝 g) (𝓝 0), - have : ∀ (i : α × β →₁[μ.prod ν] E), measurable (λ z, (nnnorm (i z - g z) : ℝ≥0∞)) := + ∫⁻ x, ∫⁻ (y : β), ∥i (x, y) - g (x, y)∥₊ ∂ν ∂μ) (𝓝 g) (𝓝 0), + have : ∀ (i : α × β →₁[μ.prod ν] E), measurable (λ z, (∥i z - g z∥₊ : ℝ≥0∞)) := λ i, ((Lp.strongly_measurable i).sub (Lp.strongly_measurable g)).ennnorm, simp_rw [← lintegral_prod_of_measurable _ (this _), ← L1.of_real_norm_sub_eq_lintegral, ← of_real_zero], diff --git a/src/measure_theory/function/conditional_expectation.lean b/src/measure_theory/function/conditional_expectation.lean index 593f0e8198d79..ebb6c10f6b9ed 100644 --- a/src/measure_theory/function/conditional_expectation.lean +++ b/src/measure_theory/function/conditional_expectation.lean @@ -1206,7 +1206,7 @@ begin of_real_integral_norm_eq_lintegral_nnnorm], swap, { rw [← mem_ℒp_one_iff_integrable], exact Lp.mem_ℒp _, }, have h_eq : ∫⁻ a, ∥condexp_ind_L1_fin hm hs hμs x a∥₊ ∂μ - = ∫⁻ a, nnnorm (condexp_ind_smul hm hs hμs x a) ∂μ, + = ∫⁻ a, ∥condexp_ind_smul hm hs hμs x a∥₊ ∂μ, { refine lintegral_congr_ae _, refine (condexp_ind_L1_fin_ae_eq_condexp_ind_smul hm hs hμs x).mono (λ z hz, _), dsimp only, diff --git a/src/measure_theory/function/l1_space.lean b/src/measure_theory/function/l1_space.lean index 807aed7299317..13dc10897967d 100644 --- a/src/measure_theory/function/l1_space.lean +++ b/src/measure_theory/function/l1_space.lean @@ -29,7 +29,7 @@ classes of integrable functions, already defined as a special case of `L^p` spac ## Main definitions * Let `f : α → β` be a function, where `α` is a `measure_space` and `β` a `normed_group`. - Then `has_finite_integral f` means `(∫⁻ a, nnnorm (f a)) < ∞`. + Then `has_finite_integral f` means `(∫⁻ a, ∥f a∥₊) < ∞`. * If `β` is moreover a `measurable_space` then `f` is called `integrable` if `f` is `measurable` and `has_finite_integral f` holds. @@ -60,7 +60,7 @@ namespace measure_theory /-! ### Some results about the Lebesgue integral involving a normed group -/ lemma lintegral_nnnorm_eq_lintegral_edist (f : α → β) : - ∫⁻ a, nnnorm (f a) ∂μ = ∫⁻ a, edist (f a) 0 ∂μ := + ∫⁻ a, ∥f a∥₊ ∂μ = ∫⁻ a, edist (f a) 0 ∂μ := by simp only [edist_eq_coe_nnnorm] lemma lintegral_norm_eq_lintegral_edist (f : α → β) : @@ -77,15 +77,15 @@ begin apply edist_triangle_right end -lemma lintegral_nnnorm_zero : ∫⁻ a : α, nnnorm (0 : β) ∂μ = 0 := by simp +lemma lintegral_nnnorm_zero : ∫⁻ a : α, ∥(0 : β)∥₊ ∂μ = 0 := by simp lemma lintegral_nnnorm_add {f : α → β} {g : α → γ} (hf : ae_strongly_measurable f μ) (hg : ae_strongly_measurable g μ) : - ∫⁻ a, nnnorm (f a) + nnnorm (g a) ∂μ = ∫⁻ a, nnnorm (f a) ∂μ + ∫⁻ a, nnnorm (g a) ∂μ := + ∫⁻ a, ∥f a∥₊ + ∥g a∥₊ ∂μ = ∫⁻ a, ∥f a∥₊ ∂μ + ∫⁻ a, ∥g a∥₊ ∂μ := lintegral_add' hf.ennnorm hg.ennnorm lemma lintegral_nnnorm_neg {f : α → β} : - ∫⁻ a, nnnorm ((-f) a) ∂μ = ∫⁻ a, nnnorm (f a) ∂μ := + ∫⁻ a, ∥(-f) a∥₊ ∂μ = ∫⁻ a, ∥f a∥₊ ∂μ := by simp only [pi.neg_apply, nnnorm_neg] /-! ### The predicate `has_finite_integral` -/ @@ -93,7 +93,7 @@ by simp only [pi.neg_apply, nnnorm_neg] /-- `has_finite_integral f μ` means that the integral `∫⁻ a, ∥f a∥ ∂μ` is finite. `has_finite_integral f` means `has_finite_integral f volume`. -/ def has_finite_integral {m : measurable_space α} (f : α → β) (μ : measure α . volume_tac) : Prop := -∫⁻ a, nnnorm (f a) ∂μ < ∞ +∫⁻ a, ∥f a∥₊ ∂μ < ∞ lemma has_finite_integral_iff_norm (f : α → β) : has_finite_integral f μ ↔ ∫⁻ a, (ennreal.of_real ∥f a∥) ∂μ < ∞ := @@ -207,7 +207,7 @@ by simpa [has_finite_integral] using hfi lemma has_finite_integral.norm {f : α → β} (hfi : has_finite_integral f μ) : has_finite_integral (λa, ∥f a∥) μ := -have eq : (λa, (nnnorm ∥f a∥ : ℝ≥0∞)) = λa, (nnnorm (f a) : ℝ≥0∞), +have eq : (λa, (nnnorm ∥f a∥ : ℝ≥0∞)) = λa, (∥f a∥₊ : ℝ≥0∞), by { funext, rw nnnorm_norm }, by { rwa [has_finite_integral, eq] } @@ -357,7 +357,7 @@ lemma has_finite_integral.smul (c : 𝕜) {f : α → β} : has_finite_integral begin simp only [has_finite_integral], assume hfi, calc - ∫⁻ (a : α), nnnorm (c • f a) ∂μ = ∫⁻ (a : α), (nnnorm c) * nnnorm (f a) ∂μ : + ∫⁻ (a : α), ∥c • f a∥₊ ∂μ = ∫⁻ (a : α), (∥c∥₊) * ∥f a∥₊ ∂μ : by simp only [nnnorm_smul, ennreal.coe_mul] ... < ∞ : begin @@ -562,7 +562,7 @@ variables {α β μ} lemma integrable.add' {f g : α → β} (hf : integrable f μ) (hg : integrable g μ) : has_finite_integral (f + g) μ := -calc ∫⁻ a, nnnorm (f a + g a) ∂μ ≤ ∫⁻ a, nnnorm (f a) + nnnorm (g a) ∂μ : +calc ∫⁻ a, ∥f a + g a∥₊ ∂μ ≤ ∫⁻ a, ∥f a∥₊ + ∥g a∥₊ ∂μ : lintegral_mono (λ a, by exact_mod_cast nnnorm_add_le _ _) ... = _ : lintegral_nnnorm_add hf.ae_strongly_measurable hg.ae_strongly_measurable ... < ∞ : add_lt_top.2 ⟨hf.has_finite_integral, hg.has_finite_integral⟩ @@ -589,7 +589,7 @@ lemma integrable.neg {f : α → β} (hf : integrable f μ) : integrable (-f) μ lemma integrable.sub' {f g : α → β} (hf : integrable f μ) (hg : integrable g μ) : has_finite_integral (f - g) μ := -calc ∫⁻ a, nnnorm (f a - g a) ∂μ ≤ ∫⁻ a, nnnorm (f a) + nnnorm (-g a) ∂μ : +calc ∫⁻ a, ∥f a - g a∥₊ ∂μ ≤ ∫⁻ a, ∥f a∥₊ + ∥-g a∥₊ ∂μ : lintegral_mono (assume a, by { simp only [sub_eq_add_neg], exact_mod_cast nnnorm_add_le _ _ } ) ... = _ : by { simp only [nnnorm_neg], @@ -949,13 +949,13 @@ variables {E : Type*} {m0 : measurable_space α} [normed_group E] lemma integrable_of_forall_fin_meas_le' {μ : measure α} (hm : m ≤ m0) [sigma_finite (μ.trim hm)] (C : ℝ≥0∞) (hC : C < ∞) {f : α → E} (hf_meas : ae_strongly_measurable f μ) - (hf : ∀ s, measurable_set[m] s → μ s ≠ ∞ → ∫⁻ x in s, nnnorm (f x) ∂μ ≤ C) : + (hf : ∀ s, measurable_set[m] s → μ s ≠ ∞ → ∫⁻ x in s, ∥f x∥₊ ∂μ ≤ C) : integrable f μ := ⟨hf_meas, (lintegral_le_of_forall_fin_meas_le' hm C hf_meas.ennnorm hf).trans_lt hC⟩ lemma integrable_of_forall_fin_meas_le [sigma_finite μ] (C : ℝ≥0∞) (hC : C < ∞) {f : α → E} (hf_meas : ae_strongly_measurable f μ) - (hf : ∀ s : set α, measurable_set s → μ s ≠ ∞ → ∫⁻ x in s, nnnorm (f x) ∂μ ≤ C) : + (hf : ∀ s : set α, measurable_set s → μ s ≠ ∞ → ∫⁻ x in s, ∥f x∥₊ ∂μ ≤ C) : integrable f μ := @integrable_of_forall_fin_meas_le' _ _ _ _ _ _ _ (by rwa trim_eq_self) C hC _ hf_meas hf @@ -1055,14 +1055,14 @@ lemma dist_def (f g : α →₁[μ] β) : by { simp [Lp.dist_def, snorm, snorm'], simp [edist_eq_coe_nnnorm_sub] } lemma norm_def (f : α →₁[μ] β) : - ∥f∥ = (∫⁻ a, nnnorm (f a) ∂μ).to_real := + ∥f∥ = (∫⁻ a, ∥f a∥₊ ∂μ).to_real := by { simp [Lp.norm_def, snorm, snorm'] } /-- Computing the norm of a difference between two L¹-functions. Note that this is not a special case of `norm_def` since `(f - g) x` and `f x - g x` are not equal (but only a.e.-equal). -/ lemma norm_sub_eq_lintegral (f g : α →₁[μ] β) : - ∥f - g∥ = (∫⁻ x, (nnnorm (f x - g x) : ℝ≥0∞) ∂μ).to_real := + ∥f - g∥ = (∫⁻ x, (∥f x - g x∥₊ : ℝ≥0∞) ∂μ).to_real := begin rw [norm_def], congr' 1, @@ -1072,14 +1072,14 @@ begin end lemma of_real_norm_eq_lintegral (f : α →₁[μ] β) : - ennreal.of_real ∥f∥ = ∫⁻ x, (nnnorm (f x) : ℝ≥0∞) ∂μ := + ennreal.of_real ∥f∥ = ∫⁻ x, (∥f x∥₊ : ℝ≥0∞) ∂μ := by { rw [norm_def, ennreal.of_real_to_real], exact ne_of_lt (has_finite_integral_coe_fn f) } /-- Computing the norm of a difference between two L¹-functions. Note that this is not a special case of `of_real_norm_eq_lintegral` since `(f - g) x` and `f x - g x` are not equal (but only a.e.-equal). -/ lemma of_real_norm_sub_eq_lintegral (f g : α →₁[μ] β) : - ennreal.of_real ∥f - g∥ = ∫⁻ x, (nnnorm (f x - g x) : ℝ≥0∞) ∂μ := + ennreal.of_real ∥f - g∥ = ∫⁻ x, (∥f x - g x∥₊ : ℝ≥0∞) ∂μ := begin simp_rw [of_real_norm_eq_lintegral, ← edist_eq_coe_nnnorm], apply lintegral_congr_ae, diff --git a/src/measure_theory/function/l2_space.lean b/src/measure_theory/function/l2_space.lean index 4c47516f2b619..b5fa74214189a 100644 --- a/src/measure_theory/function/l2_space.lean +++ b/src/measure_theory/function/l2_space.lean @@ -98,7 +98,7 @@ instance : has_inner 𝕜 (α →₂[μ] E) := ⟨λ f g, ∫ a, ⟪f a, g a⟫ lemma inner_def (f g : α →₂[μ] E) : ⟪f, g⟫ = ∫ a : α, ⟪f a, g a⟫ ∂μ := rfl lemma integral_inner_eq_sq_snorm (f : α →₂[μ] E) : - ∫ a, ⟪f a, f a⟫ ∂μ = ennreal.to_real ∫⁻ a, (nnnorm (f a) : ℝ≥0∞) ^ (2:ℝ) ∂μ := + ∫ a, ⟪f a, f a⟫ ∂μ = ennreal.to_real ∫⁻ a, (∥f a∥₊ : ℝ≥0∞) ^ (2:ℝ) ∂μ := begin simp_rw inner_self_eq_norm_sq_to_K, norm_cast, diff --git a/src/measure_theory/function/lp_space.lean b/src/measure_theory/function/lp_space.lean index 425c5a936b3fa..ad2b04f57c8b9 100644 --- a/src/measure_theory/function/lp_space.lean +++ b/src/measure_theory/function/lp_space.lean @@ -2379,7 +2379,7 @@ begin one_div_one_div p], simp_rw snorm' at hn, have h_nnnorm_nonneg : - (λ a, (nnnorm (∑ i in finset.range (n + 1), ∥f (i + 1) a - f i a∥) : ℝ≥0∞) ^ p) + (λ a, (∥∑ i in finset.range (n + 1), ∥f (i + 1) a - f i a∥∥₊ : ℝ≥0∞) ^ p) = λ a, (∑ i in finset.range (n + 1), (∥f (i + 1) a - f i a∥₊ : ℝ≥0∞)) ^ p, { ext1 a, congr, diff --git a/src/measure_theory/function/simple_func_dense_lp.lean b/src/measure_theory/function/simple_func_dense_lp.lean index db4266d1def6c..f30c330a0f56e 100644 --- a/src/measure_theory/function/simple_func_dense_lp.lean +++ b/src/measure_theory/function/simple_func_dense_lp.lean @@ -259,8 +259,8 @@ let ⟨C, hfC⟩ := f.exists_forall_norm_le in mem_ℒp_top_of_bound f.ae_strongly_measurable C $ eventually_of_forall hfC protected lemma snorm'_eq {p : ℝ} (f : α →ₛ F) (μ : measure α) : - snorm' f p μ = (∑ y in f.range, (nnnorm y : ℝ≥0∞) ^ p * μ (f ⁻¹' {y})) ^ (1/p) := -have h_map : (λ a, (nnnorm (f a) : ℝ≥0∞) ^ p) = f.map (λ a : F, (nnnorm a : ℝ≥0∞) ^ p), by simp, + snorm' f p μ = (∑ y in f.range, (∥y∥₊ : ℝ≥0∞) ^ p * μ (f ⁻¹' {y})) ^ (1/p) := +have h_map : (λ a, (∥f a∥₊ : ℝ≥0∞) ^ p) = f.map (λ a : F, (∥a∥₊ : ℝ≥0∞) ^ p), by simp, by rw [snorm', h_map, lintegral_eq_lintegral, map_lintegral] lemma measure_preimage_lt_top_of_mem_ℒp (hp_pos : p ≠ 0) (hp_ne_top : p ≠ ∞) (f : α →ₛ E) diff --git a/src/measure_theory/function/strongly_measurable.lean b/src/measure_theory/function/strongly_measurable.lean index 3f9fc7d7c5608..ac8b93e9276e1 100644 --- a/src/measure_theory/function/strongly_measurable.lean +++ b/src/measure_theory/function/strongly_measurable.lean @@ -698,12 +698,12 @@ continuous_norm.comp_strongly_measurable hf protected lemma nnnorm {m : measurable_space α} {β : Type*} [normed_group β] {f : α → β} (hf : strongly_measurable f) : - strongly_measurable (λ x, nnnorm (f x)) := + strongly_measurable (λ x, ∥f x∥₊) := continuous_nnnorm.comp_strongly_measurable hf protected lemma ennnorm {m : measurable_space α} {β : Type*} [normed_group β] {f : α → β} (hf : strongly_measurable f) : - measurable (λ a, (nnnorm (f a) : ℝ≥0∞)) := + measurable (λ a, (∥f a∥₊ : ℝ≥0∞)) := (ennreal.continuous_coe.comp_strongly_measurable hf.nnnorm).measurable protected lemma real_to_nnreal {m : measurable_space α} {f : α → ℝ} @@ -1227,11 +1227,11 @@ protected lemma norm {β : Type*} [normed_group β] {f : α → β} (hf : ae_str continuous_norm.comp_ae_strongly_measurable hf protected lemma nnnorm {β : Type*} [normed_group β] {f : α → β} (hf : ae_strongly_measurable f μ) : - ae_strongly_measurable (λ x, nnnorm (f x)) μ := + ae_strongly_measurable (λ x, ∥f x∥₊) μ := continuous_nnnorm.comp_ae_strongly_measurable hf protected lemma ennnorm {β : Type*} [normed_group β] {f : α → β} (hf : ae_strongly_measurable f μ) : - ae_measurable (λ a, (nnnorm (f a) : ℝ≥0∞)) μ := + ae_measurable (λ a, (∥f a∥₊ : ℝ≥0∞)) μ := (ennreal.continuous_coe.comp_ae_strongly_measurable hf.nnnorm).ae_measurable protected lemma edist {β : Type*} [normed_group β] {f g : α → β} diff --git a/src/measure_theory/integral/set_to_l1.lean b/src/measure_theory/integral/set_to_l1.lean index b6ce5f0f0f2aa..e98d1aac0ef20 100644 --- a/src/measure_theory/integral/set_to_l1.lean +++ b/src/measure_theory/integral/set_to_l1.lean @@ -696,7 +696,7 @@ lemma norm_eq_sum_mul (f : α →₁ₛ[μ] G) : ∥f∥ = ∑ x in (to_simple_func f).range, (μ ((to_simple_func f) ⁻¹' {x})).to_real * ∥x∥ := begin rw [norm_to_simple_func, snorm_one_eq_lintegral_nnnorm], - have h_eq := simple_func.map_apply (λ x, (nnnorm x : ℝ≥0∞)) (to_simple_func f), + have h_eq := simple_func.map_apply (λ x, (∥x∥₊ : ℝ≥0∞)) (to_simple_func f), dsimp only at h_eq, simp_rw ← h_eq, rw [simple_func.lintegral_eq_lintegral, simple_func.map_lintegral, ennreal.to_real_sum], From 2402b4d80ade4d5f7aa1a757da116368b18686b6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 4 May 2022 07:50:35 +0000 Subject: [PATCH 069/153] feat(set_theory/game/pgame): Tweak `pgame.neg` API (#13617) We modify the API for `pgame.neg` in various ways: - `left_moves_neg` and `right_moves_neg` are turned from type equivalences into type equalities. - The former equivalences are prefixed with `to_` and inverted. We also golf a few theorems. --- src/set_theory/game/impartial.lean | 25 +++---- src/set_theory/game/pgame.lean | 113 +++++++++++++++++------------ 2 files changed, 77 insertions(+), 61 deletions(-) diff --git a/src/set_theory/game/impartial.lean b/src/set_theory/game/impartial.lean index da579866c0bec..30be3c1e85ff3 100644 --- a/src/set_theory/game/impartial.lean +++ b/src/set_theory/game/impartial.lean @@ -96,16 +96,13 @@ instance impartial_neg : ∀ (G : pgame) [G.impartial], (-G).impartial begin introI hG, rw impartial_def, - split, + refine ⟨_, λ i, _, λ i, _⟩, { rw neg_neg, - symmetry, - exact neg_equiv_self G }, - split, - all_goals - { intro i, - equiv_rw G.left_moves_neg at i <|> equiv_rw G.right_moves_neg at i, - simp only [move_left_left_moves_neg_symm, move_right_right_moves_neg_symm], - exact impartial_neg _ } + exact (neg_equiv_self G).symm }, + { rw move_left_neg', + apply impartial_neg }, + { rw move_right_neg', + apply impartial_neg } end using_well_founded { dec_tac := pgame_wf_tac } @@ -196,10 +193,12 @@ lemma no_good_right_moves_iff_first_loses (G : pgame) [G.impartial] : begin rw [first_loses_of_equiv_iff (neg_equiv_self G), ←no_good_left_moves_iff_first_loses], refine ⟨λ h i, _, λ h i, _⟩, - { simpa [first_wins_of_equiv_iff (neg_equiv_self ((-G).move_left i))] - using h (left_moves_neg _ i) }, - { simpa [first_wins_of_equiv_iff (neg_equiv_self (G.move_right i))] - using h ((left_moves_neg _).symm i) } + { rw [move_left_neg', + ←first_wins_of_equiv_iff (neg_equiv_self (G.move_right (to_left_moves_neg.symm i)))], + apply h }, + { rw [move_right_neg_symm', + ←first_wins_of_equiv_iff (neg_equiv_self ((-G).move_left (to_left_moves_neg i)))], + apply h } end lemma good_left_move_iff_first_wins (G : pgame) [G.impartial] : diff --git a/src/set_theory/game/pgame.lean b/src/set_theory/game/pgame.lean index 85b789eb21f05..87d9cdc483859 100644 --- a/src/set_theory/game/pgame.lean +++ b/src/set_theory/game/pgame.lean @@ -718,34 +718,58 @@ begin exact this (list.length_map _ _).symm ha } } end -/-- An explicit equivalence between the moves for Left in `-x` and the moves for Right in `x`. -/ --- This equivalence is useful to avoid having to use `cases` unnecessarily. -def left_moves_neg (x : pgame) : (-x).left_moves ≃ x.right_moves := +theorem left_moves_neg : ∀ x : pgame, (-x).left_moves = x.right_moves +| ⟨_, _, _, _⟩ := rfl + +theorem right_moves_neg : ∀ x : pgame, (-x).right_moves = x.left_moves +| ⟨_, _, _, _⟩ := rfl + +/-- Turns a right move for `x` into a left move for `-x` and vice versa. + +Even though these types are the same (not definitionally so), this is the preferred way to convert +between them. -/ +def to_left_moves_neg {x : pgame} : x.right_moves ≃ (-x).left_moves := +equiv.cast (left_moves_neg x).symm + +/-- Turns a left move for `x` into a right move for `-x` and vice versa. + +Even though these types are the same (not definitionally so), this is the preferred way to convert +between them. -/ +def to_right_moves_neg {x : pgame} : x.left_moves ≃ (-x).right_moves := +equiv.cast (right_moves_neg x).symm + +lemma move_left_neg {x : pgame} (i) : + (-x).move_left (to_left_moves_neg i) = -x.move_right i := by { cases x, refl } -/-- An explicit equivalence between the moves for Right in `-x` and the moves for Left in `x`. -/ -def right_moves_neg (x : pgame) : (-x).right_moves ≃ x.left_moves := +@[simp] lemma move_left_neg' {x : pgame} (i) : + (-x).move_left i = -x.move_right (to_left_moves_neg.symm i) := by { cases x, refl } -@[simp] lemma move_right_left_moves_neg {x : pgame} (i : left_moves (-x)) : - move_right x ((left_moves_neg x) i) = -(move_left (-x) i) := -begin - induction x, - exact (neg_neg _).symm -end -@[simp] lemma move_left_left_moves_neg_symm {x : pgame} (i : right_moves x) : - move_left (-x) ((left_moves_neg x).symm i) = -(move_right x i) := +lemma move_right_neg {x : pgame} (i) : + (-x).move_right (to_right_moves_neg i) = -(x.move_left i) := by { cases x, refl } -@[simp] lemma move_left_right_moves_neg {x : pgame} (i : right_moves (-x)) : - move_left x ((right_moves_neg x) i) = -(move_right (-x) i) := -begin - induction x, - exact (neg_neg _).symm -end -@[simp] lemma move_right_right_moves_neg_symm {x : pgame} (i : left_moves x) : - move_right (-x) ((right_moves_neg x).symm i) = -(move_left x i) := + +@[simp] lemma move_right_neg' {x : pgame} (i) : + (-x).move_right i = -x.move_left (to_right_moves_neg.symm i) := by { cases x, refl } +lemma move_left_neg_symm {x : pgame} (i) : + x.move_left (to_right_moves_neg.symm i) = -(-x).move_right i := +by simp + +lemma move_left_neg_symm' {x : pgame} (i) : + x.move_left i = -(-x).move_right (to_right_moves_neg i) := +by simp + +lemma move_right_neg_symm {x : pgame} (i) : + x.move_right (to_left_moves_neg.symm i) = -(-x).move_left i := +by simp + +lemma move_right_neg_symm' {x : pgame} (i) : + x.move_right i = -(-x).move_left (to_left_moves_neg i) := +by simp + /-- If `x` has the same moves as `y`, then `-x` has the sames moves as `-y`. -/ def relabelling.neg_congr : ∀ {x y : pgame}, x.relabelling y → (-x).relabelling (-y) | (mk xl xr xL xR) (mk yl yr yL yR) ⟨L_equiv, R_equiv, L_relabelling, R_relabelling⟩ := @@ -756,34 +780,27 @@ def relabelling.neg_congr : ∀ {x y : pgame}, x.relabelling y → (-x).relabell theorem le_iff_neg_ge : Π {x y : pgame}, x ≤ y ↔ -y ≤ -x | (mk xl xr xL xR) (mk yl yr yL yR) := begin - rw [le_def], - rw [le_def], + rw [le_def, le_def], dsimp [neg], - split, - { intro h, - split, - { intro i, have t := h.right i, cases t, - { right, cases t, - use (@right_moves_neg (yR i)).symm t_w, convert le_iff_neg_ge.1 t_h, simp }, - { left, cases t, - use t_w, exact le_iff_neg_ge.1 t_h, } }, - { intro j, have t := h.left j, cases t, - { right, cases t, - use t_w, exact le_iff_neg_ge.1 t_h, }, - { left, cases t, - use (@left_moves_neg (xL j)).symm t_w, convert le_iff_neg_ge.1 t_h, simp, } } }, - { intro h, - split, - { intro i, have t := h.right i, cases t, - { right, cases t, - use (@left_moves_neg (xL i)) t_w, convert le_iff_neg_ge.2 _, convert t_h, simp, }, - { left, cases t, - use t_w, exact le_iff_neg_ge.2 t_h, } }, - { intro j, have t := h.left j, cases t, - { right, cases t, - use t_w, exact le_iff_neg_ge.2 t_h, }, - { left, cases t, - use (@right_moves_neg (yR j)) t_w, convert le_iff_neg_ge.2 _, convert t_h, simp } } }, + refine ⟨λ h, ⟨λ i, _, λ j, _⟩, λ h, ⟨λ i, _, λ j, _⟩⟩, + { rcases h.right i with ⟨w, h⟩ | ⟨w, h⟩, + { refine or.inr ⟨to_right_moves_neg w, _⟩, + convert le_iff_neg_ge.1 h, + rw move_right_neg }, + { exact or.inl ⟨w, le_iff_neg_ge.1 h⟩ } }, + { rcases h.left j with ⟨w, h⟩ | ⟨w, h⟩, + { exact or.inr ⟨w, le_iff_neg_ge.1 h⟩ }, + { refine or.inl ⟨to_left_moves_neg w, _⟩, + convert le_iff_neg_ge.1 h, + rw move_left_neg } }, + { rcases h.right i with ⟨w, h⟩ | ⟨w, h⟩, + { refine or.inr ⟨to_left_moves_neg.symm w, le_iff_neg_ge.2 _⟩, + rwa [move_right_neg_symm, neg_neg] }, + { exact or.inl ⟨w, le_iff_neg_ge.2 h⟩ } }, + { rcases h.left j with ⟨w, h⟩ | ⟨w, h⟩, + { exact or.inr ⟨w, le_iff_neg_ge.2 h⟩ }, + { refine or.inl ⟨to_right_moves_neg.symm w, le_iff_neg_ge.2 _⟩, + rwa [move_left_neg_symm, neg_neg] } }, end using_well_founded { dec_tac := pgame_wf_tac } From bd23639bf5754851c49f2ad05f1f8aef54bd3829 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 4 May 2022 07:50:37 +0000 Subject: [PATCH 070/153] feat(topology/bornology): add more instances (#13621) Co-authored-by: Eric Wieser --- src/topology/bornology/constructions.lean | 147 ++++++++++++++++++++++ 1 file changed, 147 insertions(+) create mode 100644 src/topology/bornology/constructions.lean diff --git a/src/topology/bornology/constructions.lean b/src/topology/bornology/constructions.lean new file mode 100644 index 0000000000000..1886ce80562f2 --- /dev/null +++ b/src/topology/bornology/constructions.lean @@ -0,0 +1,147 @@ +/- +Copyright (c) 2022 Yury G. Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury G. Kudryashov +-/ +import topology.bornology.basic + +/-! +# Bornology structure on products and subtypes + +In this file we define `bornology` and `bounded_space` instances on `α × β`, `Π i, π i`, and +`{x // p x}`. We also prove basic lemmas about `bornology.cobounded` and `bornology.is_bounded` +on these types. +-/ + +open set filter bornology function +open_locale filter + +variables {α β ι : Type*} {π : ι → Type*} [fintype ι] [bornology α] [bornology β] + [Π i, bornology (π i)] + +instance : bornology (α × β) := +{ cobounded := (cobounded α).coprod (cobounded β), + le_cofinite := @coprod_cofinite α β ▸ coprod_mono ‹bornology α›.le_cofinite + ‹bornology β›.le_cofinite } + +instance : bornology (Π i, π i) := +{ cobounded := filter.Coprod (λ i, cobounded (π i)), + le_cofinite := @Coprod_cofinite ι π _ ▸ (filter.Coprod_mono $ λ i, bornology.le_cofinite _) } + +/-- Inverse image of a bornology. -/ +@[reducible] def bornology.induced {α β : Type*} [bornology β] (f : α → β) : bornology α := +{ cobounded := comap f (cobounded β), + le_cofinite := (comap_mono (bornology.le_cofinite β)).trans (comap_cofinite_le _) } + +instance {p : α → Prop} : bornology (subtype p) := bornology.induced (coe : subtype p → α) + +namespace bornology + +/-! +### Bounded sets in `α × β` +-/ + +lemma cobounded_prod : cobounded (α × β) = (cobounded α).coprod (cobounded β) := rfl + +lemma is_bounded_image_fst_and_snd {s : set (α × β)} : + is_bounded (prod.fst '' s) ∧ is_bounded (prod.snd '' s) ↔ is_bounded s := +compl_mem_coprod.symm + +variables {s : set α} {t : set β} {S : Π i, set (π i)} + +lemma is_bounded.fst_of_prod (h : is_bounded (s ×ˢ t)) (ht : t.nonempty) : + is_bounded s := +fst_image_prod s ht ▸ (is_bounded_image_fst_and_snd.2 h).1 + +lemma is_bounded.snd_of_prod (h : is_bounded (s ×ˢ t)) (hs : s.nonempty) : + is_bounded t := +snd_image_prod hs t ▸ (is_bounded_image_fst_and_snd.2 h).2 + +lemma is_bounded.prod (hs : is_bounded s) (ht : is_bounded t) : is_bounded (s ×ˢ t) := +is_bounded_image_fst_and_snd.1 + ⟨hs.subset $ fst_image_prod_subset _ _, ht.subset $ snd_image_prod_subset _ _⟩ + +lemma is_bounded_prod_of_nonempty (hne : set.nonempty (s ×ˢ t)) : + is_bounded (s ×ˢ t) ↔ is_bounded s ∧ is_bounded t := +⟨λ h, ⟨h.fst_of_prod hne.snd, h.snd_of_prod hne.fst⟩, λ h, h.1.prod h.2⟩ + +lemma is_bounded_prod : is_bounded (s ×ˢ t) ↔ s = ∅ ∨ t = ∅ ∨ is_bounded s ∧ is_bounded t := +begin + rcases s.eq_empty_or_nonempty with rfl|hs, { simp }, + rcases t.eq_empty_or_nonempty with rfl|ht, { simp }, + simp only [hs.ne_empty, ht.ne_empty, is_bounded_prod_of_nonempty (hs.prod ht), false_or] +end + +lemma is_bounded_prod_self : is_bounded (s ×ˢ s) ↔ is_bounded s := +begin + rcases s.eq_empty_or_nonempty with rfl|hs, { simp }, + exact (is_bounded_prod_of_nonempty (hs.prod hs)).trans (and_self _) +end + +/-! +### Bounded sets in `Π i, π i` +-/ + +lemma cobounded_pi : cobounded (Π i, π i) = filter.Coprod (λ i, cobounded (π i)) := rfl + +lemma forall_is_bounded_image_eval_iff {s : set (Π i, π i)} : + (∀ i, is_bounded (eval i '' s)) ↔ is_bounded s := +compl_mem_Coprod.symm + +lemma is_bounded.pi (h : ∀ i, is_bounded (S i)) : is_bounded (pi univ S) := +forall_is_bounded_image_eval_iff.1 $ λ i, (h i).subset eval_image_univ_pi_subset + +lemma is_bounded_pi_of_nonempty (hne : (pi univ S).nonempty) : + is_bounded (pi univ S) ↔ ∀ i, is_bounded (S i) := +⟨λ H i, @eval_image_univ_pi _ _ _ i hne ▸ forall_is_bounded_image_eval_iff.2 H i, is_bounded.pi⟩ + +lemma is_bounded_pi : is_bounded (pi univ S) ↔ (∃ i, S i = ∅) ∨ ∀ i, is_bounded (S i) := +begin + by_cases hne : ∃ i, S i = ∅, + { simp [hne, univ_pi_eq_empty_iff.2 hne] }, + { simp only [hne, false_or], + simp only [not_exists, ← ne.def, ne_empty_iff_nonempty, ← univ_pi_nonempty_iff] at hne, + exact is_bounded_pi_of_nonempty hne } +end + +/-! +### Bounded sets in `{x // p x}` +-/ + +lemma is_bounded_induced {α β : Type*} [bornology β] {f : α → β} {s : set α} : + @is_bounded α (bornology.induced f) s ↔ is_bounded (f '' s) := +compl_mem_comap + +lemma is_bounded_image_subtype_coe {p : α → Prop} {s : set {x // p x}} : + is_bounded (coe '' s : set α) ↔ is_bounded s := +is_bounded_induced.symm + +end bornology + +/-! +### Bounded spaces +-/ + +open bornology + +instance [bounded_space α] [bounded_space β] : bounded_space (α × β) := +by simp [← cobounded_eq_bot_iff, cobounded_prod] + +instance [∀ i, bounded_space (π i)] : bounded_space (Π i, π i) := +by simp [← cobounded_eq_bot_iff, cobounded_pi] + +lemma bounded_space_induced_iff {α β : Type*} [bornology β] {f : α → β} : + @bounded_space α (bornology.induced f) ↔ is_bounded (range f) := +by rw [← is_bounded_univ, is_bounded_induced, image_univ] + +lemma bounded_space_subtype_iff {p : α → Prop} : bounded_space (subtype p) ↔ is_bounded {x | p x} := +by rw [bounded_space_induced_iff, subtype.range_coe_subtype] + +lemma bounded_space_coe_set_iff {s : set α} : bounded_space s ↔ is_bounded s := +bounded_space_subtype_iff + +alias bounded_space_subtype_iff ↔ _ bornology.is_bounded.bounded_space_subtype +alias bounded_space_coe_set_iff ↔ _ bornology.is_bounded.bounded_space_coe + +instance [bounded_space α] {p : α → Prop} : bounded_space (subtype p) := +(is_bounded.all {x | p x}).bounded_space_subtype From 36c5faacbd3ca2403a511ff154493f1bc1fca33f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 4 May 2022 07:50:38 +0000 Subject: [PATCH 071/153] feat(set_theory/game/pgame): Tweak `pgame.mul` API (#13651) We modify the API for `pgame.mul` in two ways: - `left_moves_mul` and `right_moves_mul` are turned from type equivalences into type equalities. - The former equivalences are prefixed with `to_` and inverted. --- src/set_theory/game/basic.lean | 46 ++++++++++++++++++++++------------ 1 file changed, 30 insertions(+), 16 deletions(-) diff --git a/src/set_theory/game/basic.lean b/src/set_theory/game/basic.lean index 073c8fc00ec5c..259245dabdb20 100644 --- a/src/set_theory/game/basic.lean +++ b/src/set_theory/game/basic.lean @@ -184,15 +184,29 @@ instance : has_mul pgame.{u} := { exact IHxr i y + IHyl j - IHxr i (yL j) } end⟩ -/-- An explicit description of the moves for Left in `x * y`. -/ -def left_moves_mul (x y : pgame) : (x * y).left_moves - ≃ x.left_moves × y.left_moves ⊕ x.right_moves × y.right_moves := -by { cases x, cases y, refl, } +theorem left_moves_mul : ∀ (x y : pgame.{u}), (x * y).left_moves + = (x.left_moves × y.left_moves ⊕ x.right_moves × y.right_moves) +| ⟨_, _, _, _⟩ ⟨_, _, _, _⟩ := rfl -/-- An explicit description of the moves for Right in `x * y`. -/ -def right_moves_mul (x y : pgame) : (x * y).right_moves - ≃ x.left_moves × y.right_moves ⊕ x.right_moves × y.left_moves := -by { cases x, cases y, refl, } +theorem right_moves_mul : ∀ (x y : pgame.{u}), (x * y).right_moves + = (x.left_moves × y.right_moves ⊕ x.right_moves × y.left_moves) +| ⟨_, _, _, _⟩ ⟨_, _, _, _⟩ := rfl + +/-- Turns two left or right moves for `x` and `y` into a left move for `x * y` and vice versa. + +Even though these types are the same (not definitionally so), this is the preferred way to convert +between them. -/ +def to_left_moves_mul {x y : pgame} : x.left_moves × y.left_moves ⊕ x.right_moves × y.right_moves + ≃ (x * y).left_moves := +equiv.cast (left_moves_mul x y).symm + +/-- Turns a left and a right move for `x` and `y` into a right move for `x * y` and vice versa. + +Even though these types are the same (not definitionally so), this is the preferred way to convert +between them. -/ +def to_right_moves_mul {x y : pgame} : x.left_moves × y.right_moves ⊕ x.right_moves × y.left_moves + ≃ (x * y).right_moves := +equiv.cast (right_moves_mul x y).symm @[simp] lemma mk_mul_move_left_inl {xl xr yl yr} {xL xR yL yR} {i j} : (mk xl xr xL xR * mk yl yr yL yR).move_left (sum.inl (i, j)) @@ -200,9 +214,9 @@ by { cases x, cases y, refl, } rfl @[simp] lemma mul_move_left_inl {x y : pgame} {i j} : - (x * y).move_left ((left_moves_mul x y).symm (sum.inl (i, j))) + (x * y).move_left (to_left_moves_mul (sum.inl (i, j))) = x.move_left i * y + x * y.move_left j - x.move_left i * y.move_left j := -by {cases x, cases y, refl} +by { cases x, cases y, refl } @[simp] lemma mk_mul_move_left_inr {xl xr yl yr} {xL xR yL yR} {i j} : (mk xl xr xL xR * mk yl yr yL yR).move_left (sum.inr (i, j)) @@ -210,9 +224,9 @@ by {cases x, cases y, refl} rfl @[simp] lemma mul_move_left_inr {x y : pgame} {i j} : - (x * y).move_left ((left_moves_mul x y).symm (sum.inr (i, j))) + (x * y).move_left (to_left_moves_mul (sum.inr (i, j))) = x.move_right i * y + x * y.move_right j - x.move_right i * y.move_right j := -by {cases x, cases y, refl} +by { cases x, cases y, refl } @[simp] lemma mk_mul_move_right_inl {xl xr yl yr} {xL xR yL yR} {i j} : (mk xl xr xL xR * mk yl yr yL yR).move_right (sum.inl (i, j)) @@ -220,9 +234,9 @@ by {cases x, cases y, refl} rfl @[simp] lemma mul_move_right_inl {x y : pgame} {i j} : - (x * y).move_right ((right_moves_mul x y).symm (sum.inl (i, j))) + (x * y).move_right (to_right_moves_mul (sum.inl (i, j))) = x.move_left i * y + x * y.move_right j - x.move_left i * y.move_right j := -by {cases x, cases y, refl} +by { cases x, cases y, refl } @[simp] lemma mk_mul_move_right_inr {xl xr yl yr} {xL xR yL yR} {i j} : (mk xl xr xL xR * mk yl yr yL yR).move_right (sum.inr (i,j)) @@ -230,9 +244,9 @@ by {cases x, cases y, refl} rfl @[simp] lemma mul_move_right_inr {x y : pgame} {i j} : - (x * y).move_right ((right_moves_mul x y).symm (sum.inr (i, j))) + (x * y).move_right (to_right_moves_mul (sum.inr (i, j))) = x.move_right i * y + x * y.move_left j - x.move_right i * y.move_left j := -by {cases x, cases y, refl} +by { cases x, cases y, refl } theorem quot_mul_comm : Π (x y : pgame.{u}), ⟦x * y⟧ = ⟦y * x⟧ | (mk xl xr xL xR) (mk yl yr yL yR) := From 319d50280e1294d6e37f2ab239f27a4c9d354f7d Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 4 May 2022 07:50:39 +0000 Subject: [PATCH 072/153] refactor(linear_algebra/*): more generalisations (#13668) Many further generalisations from `field` to `division_ring` in the linear algebra library. This PR changes some proofs; it's not just relaxing hypotheses. Co-authored-by: Scott Morrison --- src/data/finsupp/basic.lean | 9 +- src/linear_algebra/affine_space/basis.lean | 7 +- .../affine_space/finite_dimensional.lean | 146 ++++++------ src/linear_algebra/finite_dimensional.lean | 211 +++++++++++++----- 4 files changed, 231 insertions(+), 142 deletions(-) diff --git a/src/data/finsupp/basic.lean b/src/data/finsupp/basic.lean index 7ced9f872c9ba..473f2a0a7af0f 100644 --- a/src/data/finsupp/basic.lean +++ b/src/data/finsupp/basic.lean @@ -195,6 +195,12 @@ by simp only [set.subset_def, mem_coe, mem_support_iff]; equiv_fun_on_fintype.symm f = f := by { ext, simp [equiv_fun_on_fintype], } +/-- If `α` has a unique term, +then the type of finitely supported functions `α →₀ β` is equivalent to `β`. -/ +@[simps] noncomputable +def _root_.equiv.finsupp_unique {ι : Type*} [unique ι] : (ι →₀ M) ≃ M := +finsupp.equiv_fun_on_fintype.trans (equiv.fun_unique ι M) + end basic /-! ### Declarations about `single` -/ @@ -334,10 +340,11 @@ end lemma unique_single [unique α] (x : α →₀ M) : x = single default (x default) := ext $ unique.forall_iff.2 single_eq_same.symm +@[ext] lemma unique_ext [unique α] {f g : α →₀ M} (h : f default = g default) : f = g := ext $ λ a, by rwa [unique.eq_default a] -lemma unique_ext_iff [unique α] {f g : α →₀ M} : f = g ↔ f default = g default := +lemma unique_ext_iff [unique α] {f g : α →₀ M} : f = g ↔ f default = g default := ⟨λ h, h ▸ rfl, unique_ext⟩ @[simp] lemma unique_single_eq_iff [unique α] {b' : M} : diff --git a/src/linear_algebra/affine_space/basis.lean b/src/linear_algebra/affine_space/basis.lean index 20402c7e4bc00..f599bb429dfeb 100644 --- a/src/linear_algebra/affine_space/basis.lean +++ b/src/linear_algebra/affine_space/basis.lean @@ -357,10 +357,9 @@ end end comm_ring -section field +section division_ring --- TODO Relax `field` to `division_ring` (results are still true) -variables [field k] [module k V] +variables [division_ring k] [module k V] include V variables (k V P) @@ -388,6 +387,6 @@ begin simp [h_tot], end -end field +end division_ring end affine_basis diff --git a/src/linear_algebra/affine_space/finite_dimensional.lean b/src/linear_algebra/affine_space/finite_dimensional.lean index 2f297c4804a24..f782f009cb68a 100644 --- a/src/linear_algebra/affine_space/finite_dimensional.lean +++ b/src/linear_algebra/affine_space/finite_dimensional.lean @@ -30,7 +30,6 @@ include V open affine_subspace finite_dimensional module -section variables [division_ring k] [add_comm_group V] [module k V] [affine_space V P] /-- The `vector_span` of a finite set is finite-dimensional. -/ @@ -84,11 +83,6 @@ lemma finite_of_fin_dim_affine_independent [finite_dimensional k V] {s : set P} (hi : affine_independent k (coe : s → P)) : s.finite := ⟨fintype_of_fin_dim_affine_independent k hi⟩ -end - -section - -variables [field k] [add_comm_group V] [module k V] [affine_space V P] variables {k} /-- The `vector_span` of a finite subset of an affinely independent @@ -125,53 +119,6 @@ begin exact hi.finrank_vector_span_image_finset hc end -/-- If the `vector_span` of a finite subset of an affinely independent -family lies in a submodule with dimension one less than its -cardinality, it equals that submodule. -/ -lemma affine_independent.vector_span_image_finset_eq_of_le_of_card_eq_finrank_add_one - {p : ι → P} (hi : affine_independent k p) {s : finset ι} {sm : submodule k V} - [finite_dimensional k sm] (hle : vector_span k (s.image p : set P) ≤ sm) - (hc : finset.card s = finrank k sm + 1) : vector_span k (s.image p : set P) = sm := -eq_of_le_of_finrank_eq hle $ hi.finrank_vector_span_image_finset hc - -/-- If the `vector_span` of a finite affinely independent -family lies in a submodule with dimension one less than its -cardinality, it equals that submodule. -/ -lemma affine_independent.vector_span_eq_of_le_of_card_eq_finrank_add_one [fintype ι] - {p : ι → P} (hi : affine_independent k p) {sm : submodule k V} [finite_dimensional k sm] - (hle : vector_span k (set.range p) ≤ sm) (hc : fintype.card ι = finrank k sm + 1) : - vector_span k (set.range p) = sm := -eq_of_le_of_finrank_eq hle $ hi.finrank_vector_span hc - -/-- If the `affine_span` of a finite subset of an affinely independent -family lies in an affine subspace whose direction has dimension one -less than its cardinality, it equals that subspace. -/ -lemma affine_independent.affine_span_image_finset_eq_of_le_of_card_eq_finrank_add_one - {p : ι → P} (hi : affine_independent k p) {s : finset ι} {sp : affine_subspace k P} - [finite_dimensional k sp.direction] (hle : affine_span k (s.image p : set P) ≤ sp) - (hc : finset.card s = finrank k sp.direction + 1) : affine_span k (s.image p : set P) = sp := -begin - have hn : (s.image p).nonempty, - { rw [finset.nonempty.image_iff, ← finset.card_pos, hc], apply nat.succ_pos }, - refine eq_of_direction_eq_of_nonempty_of_le _ ((affine_span_nonempty k _).2 hn) hle, - have hd := direction_le hle, - rw direction_affine_span at ⊢ hd, - exact hi.vector_span_image_finset_eq_of_le_of_card_eq_finrank_add_one hd hc -end - -/-- If the `affine_span` of a finite affinely independent family lies -in an affine subspace whose direction has dimension one less than its -cardinality, it equals that subspace. -/ -lemma affine_independent.affine_span_eq_of_le_of_card_eq_finrank_add_one [fintype ι] - {p : ι → P} (hi : affine_independent k p) {sp : affine_subspace k P} - [finite_dimensional k sp.direction] (hle : affine_span k (set.range p) ≤ sp) - (hc : fintype.card ι = finrank k sp.direction + 1) : affine_span k (set.range p) = sp := -begin - rw ←finset.card_univ at hc, - rw [←set.image_univ, ←finset.coe_univ, ← finset.coe_image] at ⊢ hle, - exact hi.affine_span_image_finset_eq_of_le_of_card_eq_finrank_add_one hle hc -end - /-- The `vector_span` of a finite affinely independent family whose cardinality is one more than that of the finite-dimensional space is `⊤`. -/ @@ -180,24 +127,6 @@ lemma affine_independent.vector_span_eq_top_of_card_eq_finrank_add_one [finite_d vector_span k (set.range p) = ⊤ := eq_top_of_finrank_eq $ hi.finrank_vector_span hc -/-- The `affine_span` of a finite affinely independent family is `⊤` iff the -family's cardinality is one more than that of the finite-dimensional space. -/ -lemma affine_independent.affine_span_eq_top_iff_card_eq_finrank_add_one [finite_dimensional k V] - [fintype ι] {p : ι → P} (hi : affine_independent k p) : - affine_span k (set.range p) = ⊤ ↔ fintype.card ι = finrank k V + 1 := -begin - split, - { intros h_tot, - let n := fintype.card ι - 1, - have hn : fintype.card ι = n + 1, - { exact (nat.succ_pred_eq_of_pos (card_pos_of_affine_span_eq_top k V P h_tot)).symm, }, - rw [hn, ← finrank_top, ← (vector_span_eq_top_of_affine_span_eq_top k V P) h_tot, - ← hi.finrank_vector_span hn], }, - { intros hc, - rw [← finrank_top, ← direction_top k V P] at hc, - exact hi.affine_span_eq_of_le_of_card_eq_finrank_add_one le_top hc, }, -end - variables (k) /-- The `vector_span` of `n + 1` points in an indexed family has @@ -269,10 +198,74 @@ lemma finrank_vector_span_le_iff_not_affine_independent [fintype ι] (p : ι → finrank k (vector_span k (set.range p)) ≤ n ↔ ¬ affine_independent k p := (not_iff_comm.1 (affine_independent_iff_not_finrank_vector_span_le k p hc).symm).symm +variables {k} + +/-- If the `vector_span` of a finite subset of an affinely independent +family lies in a submodule with dimension one less than its +cardinality, it equals that submodule. -/ +lemma affine_independent.vector_span_image_finset_eq_of_le_of_card_eq_finrank_add_one + {p : ι → P} (hi : affine_independent k p) {s : finset ι} {sm : submodule k V} + [finite_dimensional k sm] (hle : vector_span k (s.image p : set P) ≤ sm) + (hc : finset.card s = finrank k sm + 1) : vector_span k (s.image p : set P) = sm := +eq_of_le_of_finrank_eq hle $ hi.finrank_vector_span_image_finset hc + +/-- If the `vector_span` of a finite affinely independent +family lies in a submodule with dimension one less than its +cardinality, it equals that submodule. -/ +lemma affine_independent.vector_span_eq_of_le_of_card_eq_finrank_add_one [fintype ι] + {p : ι → P} (hi : affine_independent k p) {sm : submodule k V} [finite_dimensional k sm] + (hle : vector_span k (set.range p) ≤ sm) (hc : fintype.card ι = finrank k sm + 1) : + vector_span k (set.range p) = sm := +eq_of_le_of_finrank_eq hle $ hi.finrank_vector_span hc + +/-- If the `affine_span` of a finite subset of an affinely independent +family lies in an affine subspace whose direction has dimension one +less than its cardinality, it equals that subspace. -/ +lemma affine_independent.affine_span_image_finset_eq_of_le_of_card_eq_finrank_add_one + {p : ι → P} (hi : affine_independent k p) {s : finset ι} {sp : affine_subspace k P} + [finite_dimensional k sp.direction] (hle : affine_span k (s.image p : set P) ≤ sp) + (hc : finset.card s = finrank k sp.direction + 1) : affine_span k (s.image p : set P) = sp := +begin + have hn : (s.image p).nonempty, + { rw [finset.nonempty.image_iff, ← finset.card_pos, hc], apply nat.succ_pos }, + refine eq_of_direction_eq_of_nonempty_of_le _ ((affine_span_nonempty k _).2 hn) hle, + have hd := direction_le hle, + rw direction_affine_span at ⊢ hd, + exact hi.vector_span_image_finset_eq_of_le_of_card_eq_finrank_add_one hd hc end -section division_ring -variables [division_ring k] [add_comm_group V] [module k V] [affine_space V P] +/-- If the `affine_span` of a finite affinely independent family lies +in an affine subspace whose direction has dimension one less than its +cardinality, it equals that subspace. -/ +lemma affine_independent.affine_span_eq_of_le_of_card_eq_finrank_add_one [fintype ι] + {p : ι → P} (hi : affine_independent k p) {sp : affine_subspace k P} + [finite_dimensional k sp.direction] (hle : affine_span k (set.range p) ≤ sp) + (hc : fintype.card ι = finrank k sp.direction + 1) : affine_span k (set.range p) = sp := +begin + rw ←finset.card_univ at hc, + rw [←set.image_univ, ←finset.coe_univ, ← finset.coe_image] at ⊢ hle, + exact hi.affine_span_image_finset_eq_of_le_of_card_eq_finrank_add_one hle hc +end + +/-- The `affine_span` of a finite affinely independent family is `⊤` iff the +family's cardinality is one more than that of the finite-dimensional space. -/ +lemma affine_independent.affine_span_eq_top_iff_card_eq_finrank_add_one [finite_dimensional k V] + [fintype ι] {p : ι → P} (hi : affine_independent k p) : + affine_span k (set.range p) = ⊤ ↔ fintype.card ι = finrank k V + 1 := +begin + split, + { intros h_tot, + let n := fintype.card ι - 1, + have hn : fintype.card ι = n + 1, + { exact (nat.succ_pred_eq_of_pos (card_pos_of_affine_span_eq_top k V P h_tot)).symm, }, + rw [hn, ← finrank_top, ← (vector_span_eq_top_of_affine_span_eq_top k V P) h_tot, + ← hi.finrank_vector_span hn], }, + { intros hc, + rw [← finrank_top, ← direction_top k V P] at hc, + exact hi.affine_span_eq_of_le_of_card_eq_finrank_add_one le_top hc, }, +end + +variables (k) /-- A set of points is collinear if their `vector_span` has dimension at most `1`. -/ @@ -376,11 +369,6 @@ begin simp [hp] } end -end division_ring - -section field -variables [field k] [add_comm_group V] [module k V] [affine_space V P] - /-- Three points are affinely independent if and only if they are not collinear. -/ lemma affine_independent_iff_not_collinear (p : fin 3 → P) : @@ -395,6 +383,4 @@ lemma collinear_iff_not_affine_independent (p : fin 3 → P) : by rw [collinear_iff_finrank_le_one, finrank_vector_span_le_iff_not_affine_independent k p (fintype.card_fin 3)] -end field - end affine_space' diff --git a/src/linear_algebra/finite_dimensional.lean b/src/linear_algebra/finite_dimensional.lean index 2ee8800d2efe6..bfb42092bc22c 100644 --- a/src/linear_algebra/finite_dimensional.lean +++ b/src/linear_algebra/finite_dimensional.lean @@ -82,13 +82,15 @@ Use `finite_dimensional.of_fintype_basis` to prove finite dimension from another @[reducible] def finite_dimensional (K V : Type*) [division_ring K] [add_comm_group V] [module K V] := module.finite K V +variables {K : Type u} {V : Type v} + namespace finite_dimensional open is_noetherian section division_ring -variables {K : Type u} {V : Type v} [division_ring K] [add_comm_group V] [module K V] +variables [division_ring K] [add_comm_group V] [module K V] {V₂ : Type v'} [add_comm_group V₂] [module K V₂] /-- If the codomain of an injective linear map is finite dimensional, the domain must be as well. -/ @@ -597,45 +599,52 @@ end end -end division_ring - -section field - -variables {K : Type u} {V : Type v} [field K] [add_comm_group V] [module K V] -{V₂ : Type v'} [add_comm_group V₂] [module K V₂] - /-- In a vector space with dimension 1, each set {v} is a basis for `v ≠ 0`. -/ +@[simps] noncomputable def basis_singleton (ι : Type*) [unique ι] (h : finrank K V = 1) (v : V) (hv : v ≠ 0) : basis ι K V := let b := basis_unique ι h in -b.map (linear_equiv.smul_of_unit (units.mk0 - (b.repr v default) - (mt basis_unique.repr_eq_zero_iff.mp hv))) +let h : b.repr v default ≠ 0 := mt basis_unique.repr_eq_zero_iff.mp hv in +basis.of_repr +{ to_fun := λ w, finsupp.single default (b.repr w default / b.repr v default), + inv_fun := λ f, f default • v, + map_add' := by simp [add_div], + map_smul' := by simp [mul_div], + left_inv := λ w, begin + apply_fun b.repr using b.repr.to_equiv.injective, + apply_fun equiv.finsupp_unique, + simp only [linear_equiv.map_smulₛₗ, finsupp.coe_smul, finsupp.single_eq_same, ring_hom.id_apply, + smul_eq_mul, pi.smul_apply, equiv.finsupp_unique_apply], + exact div_mul_cancel _ h, + end , + right_inv := λ f, begin + ext, + simp only [linear_equiv.map_smulₛₗ, finsupp.coe_smul, finsupp.single_eq_same, ring_hom.id_apply, + smul_eq_mul, pi.smul_apply], + exact mul_div_cancel _ h, + end, } @[simp] lemma basis_singleton_apply (ι : Type*) [unique ι] (h : finrank K V = 1) (v : V) (hv : v ≠ 0) (i : ι) : basis_singleton ι h v hv i = v := -calc basis_singleton ι h v hv i - = (((basis_unique ι h).repr) v) default • (basis_unique ι h) default : - by simp [subsingleton.elim i default, basis_singleton, linear_equiv.smul_of_unit, - units.smul_def] -... = v : by rw [← finsupp.total_unique K (basis.repr _ v), basis.total_repr] +by { cases unique.uniq ‹unique ι› i, simp [basis_singleton], } @[simp] lemma range_basis_singleton (ι : Type*) [unique ι] (h : finrank K V = 1) (v : V) (hv : v ≠ 0) : set.range (basis_singleton ι h v hv) = {v} := by rw [set.range_unique, basis_singleton_apply] -end field +end division_ring end finite_dimensional -variables {K : Type u} {V : Type v} [field K] [add_comm_group V] [module K V] -{V₂ : Type v'} [add_comm_group V₂] [module K V₂] +variables {K V} section zero_dim +variables [division_ring K] [add_comm_group V] [module K V] + open finite_dimensional lemma finite_dimensional_of_dim_eq_zero (h : module.rank K V = 0) : finite_dimensional K V := @@ -714,6 +723,9 @@ end zero_dim namespace submodule open is_noetherian finite_dimensional +section division_ring +variables [division_ring K] [add_comm_group V] [module K V] + /-- A submodule is finitely generated if and only if it is finite-dimensional -/ theorem fg_iff_finite_dimensional (s : submodule K V) : s.fg ↔ finite_dimensional K s := @@ -786,6 +798,22 @@ begin apply_instance }, end +/-- The dimension of a submodule is bounded by the dimension of the ambient space. -/ +lemma finrank_le [finite_dimensional K V] (s : submodule K V) : finrank K s ≤ finrank K V := +by simpa only [cardinal.nat_cast_le, ←finrank_eq_dim] using + s.subtype.dim_le_of_injective (injective_subtype s) + +/-- The dimension of a quotient is bounded by the dimension of the ambient space. -/ +lemma finrank_quotient_le [finite_dimensional K V] (s : submodule K V) : + finrank K (V ⧸ s) ≤ finrank K V := +by simpa only [cardinal.nat_cast_le, ←finrank_eq_dim] using + (mkq s).dim_le_of_surjective (surjective_quot_mk _) + +end division_ring + +section field +variables [field K] [add_comm_group V] [module K V] + /-- In a finite-dimensional vector space, the dimensions of a submodule and of the corresponding quotient add up to the dimension of the space. -/ theorem finrank_quotient_add_finrank [finite_dimensional K V] (s : submodule K V) : @@ -796,10 +824,6 @@ begin exact_mod_cast this end -/-- The dimension of a submodule is bounded by the dimension of the ambient space. -/ -lemma finrank_le [finite_dimensional K V] (s : submodule K V) : finrank K s ≤ finrank K V := -by { rw ← s.finrank_quotient_add_finrank, exact nat.le_add_left _ _ } - /-- The dimension of a strict submodule is strictly bounded by the dimension of the ambient space. -/ lemma finrank_lt [finite_dimensional K V] {s : submodule K V} (h : s < ⊤) : @@ -809,11 +833,6 @@ begin exact nat.lt_add_of_zero_lt_left _ _ (finrank_pos_iff.mpr (quotient.nontrivial_of_lt_top _ h)) end -/-- The dimension of a quotient is bounded by the dimension of the ambient space. -/ -lemma finrank_quotient_le [finite_dimensional K V] (s : submodule K V) : - finrank K (V ⧸ s) ≤ finrank K V := -by { rw ← s.finrank_quotient_add_finrank, exact nat.le_add_right _ _ } - /-- The sum of the dimensions of s + t and s ∩ t is the sum of the dimensions of s and t -/ theorem dim_sup_add_dim_inf_eq (s t : submodule K V) [finite_dimensional K s] [finite_dimensional K t] : @@ -840,11 +859,16 @@ begin refl, end +end field + end submodule namespace linear_equiv open finite_dimensional +variables [division_ring K] [add_comm_group V] [module K V] +{V₂ : Type v'} [add_comm_group V₂] [module K V₂] + /-- Finite dimensionality is preserved under linear equivalence. -/ protected theorem finite_dimensional (f : V ≃ₗ[K] V₂) [finite_dimensional K V] : finite_dimensional K V₂ := @@ -864,6 +888,9 @@ lemma finrank_map_eq (f : M ≃ₗ[R] M₂) (p : submodule R M) : end linear_equiv +section +variables [division_ring K] [add_comm_group V] [module K V] + instance finite_dimensional_finsupp {ι : Type*} [fintype ι] [h : finite_dimensional K V] : finite_dimensional K (ι →₀ V) := begin @@ -871,8 +898,14 @@ begin exact (finsupp.linear_equiv_fun_on_fintype K V ι).symm.finite_dimensional end +end + namespace finite_dimensional +section division_ring +variables [division_ring K] [add_comm_group V] [module K V] +{V₂ : Type v'} [add_comm_group V₂] [module K V₂] + /-- Two finite-dimensional vector spaces are isomorphic if they have the same (finite) dimension. -/ @@ -888,8 +921,6 @@ theorem nonempty_linear_equiv_iff_finrank_eq [finite_dimensional K V] [finite_di nonempty (V ≃ₗ[K] V₂) ↔ finrank K V = finrank K V₂ := ⟨λ ⟨h⟩, h.finrank_eq, λ h, nonempty_linear_equiv_of_finrank_eq h⟩ -section - variables (V V₂) /-- @@ -899,7 +930,7 @@ noncomputable def linear_equiv.of_finrank_eq [finite_dimensional K V] [finite_di (cond : finrank K V = finrank K V₂) : V ≃ₗ[K] V₂ := classical.choice $ nonempty_linear_equiv_of_finrank_eq cond -end +variables {V} lemma eq_of_le_of_finrank_le {S₁ S₂ : submodule K V} [finite_dimensional K S₂] (hle : S₁ ≤ S₂) (hd : finrank K S₂ ≤ finrank K S₁) : S₁ = S₂ := @@ -916,10 +947,16 @@ lemma eq_of_le_of_finrank_eq {S₁ S₂ : submodule K V} [finite_dimensional K S eq_of_le_of_finrank_le hle hd.ge @[simp] -lemma finrank_map_subtype_eq (p : subspace K V) (q : subspace K p) : +lemma finrank_map_subtype_eq (p : submodule K V) (q : submodule K p) : finite_dimensional.finrank K (q.map p.subtype) = finite_dimensional.finrank K q := (submodule.equiv_subtype_map p q).symm.finrank_eq +end division_ring + +section field +variables [field K] [add_comm_group V] [module K V] +{V₂ : Type v'} [add_comm_group V₂] [module K V₂] + variables [finite_dimensional K V] [finite_dimensional K V₂] /-- Given isomorphic subspaces `p q` of vector spaces `V` and `V₁` respectively, @@ -943,11 +980,17 @@ begin ← linear_equiv.finrank_eq f, add_comm, submodule.finrank_quotient_add_finrank] end +end field + end finite_dimensional namespace linear_map open finite_dimensional +section division_ring +variables [division_ring K] [add_comm_group V] [module K V] +{V₂ : Type v'} [add_comm_group V₂] [module K V₂] + /-- On a finite-dimensional space, an injective linear map is surjective. -/ lemma surjective_of_injective [finite_dimensional K V] {f : V →ₗ[K] V} (hinj : injective f) : surjective f := @@ -957,6 +1000,27 @@ begin exact range_eq_top.1 (eq_top_of_finrank_eq h.symm) end +/-- The image under an onto linear map of a finite-dimensional space is also finite-dimensional. -/ +lemma finite_dimensional_of_surjective [h : finite_dimensional K V] + (f : V →ₗ[K] V₂) (hf : f.range = ⊤) : finite_dimensional K V₂ := +module.finite.of_surjective f $ range_eq_top.1 hf + +/-- The range of a linear map defined on a finite-dimensional space is also finite-dimensional. -/ +instance finite_dimensional_range [h : finite_dimensional K V] (f : V →ₗ[K] V₂) : + finite_dimensional K f.range := +f.quot_ker_equiv_range.finite_dimensional + +/-- The dimensions of the domain and range of an injective linear map are equal. -/ +lemma finrank_range_of_inj {f : V →ₗ[K] V₂} (hf : function.injective f) : + finrank K f.range = finrank K V := +by rw (linear_equiv.of_injective f hf).finrank_eq + +end division_ring + +section field +variables [field K] [add_comm_group V] [module K V] +{V₂ : Type v'} [add_comm_group V₂] [module K V₂] + /-- On a finite-dimensional space, a linear map is injective if and only if it is surjective. -/ lemma injective_iff_surjective [finite_dimensional K V] {f : V →ₗ[K] V} : injective f ↔ surjective f := @@ -991,31 +1055,19 @@ they are inverse to each other on the other side. -/ lemma comp_eq_id_comm [finite_dimensional K V] {f g : V →ₗ[K] V} : f.comp g = id ↔ g.comp f = id := mul_eq_one_comm -/-- The image under an onto linear map of a finite-dimensional space is also finite-dimensional. -/ -lemma finite_dimensional_of_surjective [h : finite_dimensional K V] - (f : V →ₗ[K] V₂) (hf : f.range = ⊤) : finite_dimensional K V₂ := -module.finite.of_surjective f $ range_eq_top.1 hf - -/-- The range of a linear map defined on a finite-dimensional space is also finite-dimensional. -/ -instance finite_dimensional_range [h : finite_dimensional K V] (f : V →ₗ[K] V₂) : - finite_dimensional K f.range := -f.quot_ker_equiv_range.finite_dimensional - /-- rank-nullity theorem : the dimensions of the kernel and the range of a linear map add up to the dimension of the source space. -/ theorem finrank_range_add_finrank_ker [finite_dimensional K V] (f : V →ₗ[K] V₂) : finrank K f.range + finrank K f.ker = finrank K V := by { rw [← f.quot_ker_equiv_range.finrank_eq], exact submodule.finrank_quotient_add_finrank _ } -/-- The dimensions of the domain and range of an injective linear map are equal. -/ -lemma finrank_range_of_inj {f : V →ₗ[K] V₂} (hf : function.injective f) : - finrank K f.range = finrank K V := -by rw (linear_equiv.of_injective f hf).finrank_eq - +end field end linear_map namespace linear_equiv open finite_dimensional + +variables [field K] [add_comm_group V] [module K V] variables [finite_dimensional K V] /-- The linear equivalence corresponging to an injective endomorphism. -/ @@ -1037,6 +1089,8 @@ end linear_equiv namespace linear_map +variables [field K] [add_comm_group V] [module K V] + lemma is_unit_iff_ker_eq_bot [finite_dimensional K V] (f : V →ₗ[K] V): is_unit f ↔ f.ker = ⊥ := begin split, @@ -1055,6 +1109,9 @@ end linear_map open module finite_dimensional +section +variables [division_ring K] [add_comm_group V] [module K V] + section top @[simp] @@ -1076,8 +1133,13 @@ begin exact basis.empty _ end +end + namespace linear_map +variables [field K] [add_comm_group V] [module K V] +{V₂ : Type v'} [add_comm_group V₂] [module K V₂] + theorem injective_iff_surjective_of_finrank_eq_finrank [finite_dimensional K V] [finite_dimensional K V₂] (H : finrank K V = finrank K V₂) {f : V →ₗ[K] V₂} : function.injective f ↔ function.surjective f := @@ -1160,12 +1222,9 @@ end namespace submodule -lemma finrank_mono [finite_dimensional K V] : - monotone (λ (s : submodule K V), finrank K s) := -λ s t hst, -calc finrank K s = finrank K (comap t.subtype s) - : linear_equiv.finrank_eq (comap_subtype_equiv_of_le hst).symm -... ≤ finrank K t : submodule.finrank_le _ +section division_ring +variables [division_ring K] [add_comm_group V] [module K V] +{V₂ : Type v'} [add_comm_group V₂] [module K V₂] lemma lt_of_le_of_finrank_lt_finrank {s t : submodule K V} (le : s ≤ t) (lt : finrank K s < finrank K t) : s < t := @@ -1178,6 +1237,19 @@ begin exact lt_of_le_of_finrank_lt_finrank le_top lt end +lemma finrank_mono [finite_dimensional K V] : + monotone (λ (s : submodule K V), finrank K s) := +λ s t hst, +calc finrank K s = finrank K (comap t.subtype s) + : linear_equiv.finrank_eq (comap_subtype_equiv_of_le hst).symm +... ≤ finrank K t : submodule.finrank_le _ + +end division_ring + +section field +variables [field K] [add_comm_group V] [module K V] +{V₂ : Type v'} [add_comm_group V₂] [module K V₂] + lemma finrank_lt_finrank_of_lt [finite_dimensional K V] {s t : submodule K V} (hst : s < t) : finrank K s < finrank K t := begin @@ -1197,23 +1269,24 @@ begin exact finrank_top end +end field + end submodule section span open submodule -variable (K) +section division_ring +variables [division_ring K] [add_comm_group V] [module K V] +variable (K) /-- The rank of a set of vectors as a natural number. -/ protected noncomputable def set.finrank (s : set V) : ℕ := finrank K (span K s) variable {K} -lemma set.finrank_mono [finite_dimensional K V] {s t : set V} (h : s ⊆ t) : - s.finrank K ≤ t.finrank K := finrank_mono (span_mono h) - lemma finrank_span_le_card (s : set V) [fin : fintype s] : finrank K (span K s) ≤ s.to_finset.card := begin @@ -1276,10 +1349,23 @@ begin simp [hv] } end +end division_ring + +section field +variables [field K] [add_comm_group V] [module K V] + +lemma set.finrank_mono [finite_dimensional K V] {s t : set V} (h : s ⊆ t) : + s.finrank K ≤ t.finrank K := finrank_mono (span_mono h) + +end field + end span section basis +section division_ring +variables [division_ring K] [add_comm_group V] [module K V] + lemma linear_independent_of_span_eq_top_of_card_eq_finrank {ι : Type*} [fintype ι] {b : ι → V} (span_eq : span K (set.range b) = ⊤) (card_eq : fintype.card ι = finrank K V) : linear_independent K b := @@ -1381,6 +1467,11 @@ basis_of_span_eq_top_of_card_eq_finrank (coe : s → V) ((@subtype.range_coe_subtype _ s).symm ▸ span_eq) (trans s.to_finset_card.symm card_eq) +end division_ring + +section field +variables [field K] [add_comm_group V] [module K V] + lemma span_eq_top_of_linear_independent_of_card_eq_finrank {ι : Type*} [hι : nonempty ι] [fintype ι] {b : ι → V} (lin_ind : linear_independent K b) (card_eq : fintype.card ι = finrank K V) : @@ -1445,6 +1536,8 @@ basis_of_linear_independent_of_card_eq_finrank lin_ind (trans s.to_finset_card.s ⇑(set_basis_of_linear_independent_of_card_eq_finrank lin_ind card_eq) = coe := basis.coe_mk _ _ +end field + end basis /-! @@ -1452,6 +1545,8 @@ We now give characterisations of `finrank K V = 1` and `finrank K V ≤ 1`. -/ section finrank_eq_one +variables [division_ring K] [add_comm_group V] [module K V] + /-- If there is a nonzero vector and every other vector is a multiple of it, then the module has dimension one. -/ lemma finrank_eq_one (v : V) (n : v ≠ 0) (h : ∀ w : V, ∃ c : K, c • v = w) : @@ -1639,6 +1734,8 @@ end subalgebra_dim namespace module namespace End +variables [field K] [add_comm_group V] [module K V] + lemma exists_ker_pow_eq_ker_pow_succ [finite_dimensional K V] (f : End K V) : ∃ (k : ℕ), k ≤ finrank K V ∧ (f ^ k).ker = (f ^ k.succ).ker := begin From 51d816755aac557749ca7339fb4eb77e3e8996ec Mon Sep 17 00:00:00 2001 From: Aaron Anderson Date: Wed, 4 May 2022 07:50:40 +0000 Subject: [PATCH 073/153] feat(model_theory/elementary_maps): The elementary diagram of a structure (#13724) Defines the elementary diagram of a structure - the theory consisting of all sentences with parameters it satisfies. Defines the canonical elementary embedding of a structure into any model of its elementary diagram. --- src/model_theory/elementary_maps.lean | 28 +++++++++++++++++++++++++++ src/model_theory/language_map.lean | 10 ++++++++++ 2 files changed, 38 insertions(+) diff --git a/src/model_theory/elementary_maps.lean b/src/model_theory/elementary_maps.lean index 1516f47a25dd8..3a32fae56b587 100644 --- a/src/model_theory/elementary_maps.lean +++ b/src/model_theory/elementary_maps.lean @@ -14,6 +14,10 @@ import model_theory.substructures realizations of formulas. * A `first_order.language.elementary_substructure` is a substructure where the realization of each formula agrees with the realization in the larger model. +* The `first_order.language.elementary_diagram` of a structure is the set of all sentences with + parameters that the structure satisfies. +* `first_order.language.elementary_embedding.of_models_elementary_diagram` is the canonical +elementary embedding of any structure into a model of its elementary diagram. ## Main Results * The Tarski-Vaught Test for embeddings: `first_order.language.embedding.is_elementary_of_exists` @@ -156,6 +160,30 @@ lemma comp_assoc (f : M ↪ₑ[L] N) (g : N ↪ₑ[L] P) (h : P ↪ₑ[L] Q) : end elementary_embedding +variables (L) (M) + +/-- The elementary diagram of an `L`-structure is the set of all sentences with parameters it + satisfies. -/ +abbreviation elementary_diagram : L[[M]].Theory := L[[M]].complete_theory M + +/-- The canonical elementary embedding of an `L`-structure into any model of its elementary diagram +-/ +@[simps] def elementary_embedding.of_models_elementary_diagram + (N : Type*) [L.Structure N] [L[[M]].Structure N] + [(Lhom_with_constants L M).is_expansion_on N] [N ⊨ L.elementary_diagram M] : + M ↪ₑ[L] N := +⟨(coe : L[[M]].constants → N) ∘ sum.inr, λ n φ x, begin + refine trans _ ((realize_iff_of_model_complete_theory M N (((L.Lhom_with_constants + M).on_bounded_formula φ).subst (constants.term ∘ sum.inr ∘ x)).alls).trans _), + { simp_rw [sentence.realize, bounded_formula.realize_alls, bounded_formula.realize_subst, + Lhom.realize_on_bounded_formula, formula.realize, unique.forall_iff, realize_constants] }, + { simp_rw [sentence.realize, bounded_formula.realize_alls, bounded_formula.realize_subst, + Lhom.realize_on_bounded_formula, formula.realize, unique.forall_iff], + refl } +end⟩ + +variables {L M} + namespace embedding /-- The Tarski-Vaught test for elementarity of an embedding. -/ diff --git a/src/model_theory/language_map.lean b/src/model_theory/language_map.lean index 966687335630f..7c3651165c842 100644 --- a/src/model_theory/language_map.lean +++ b/src/model_theory/language_map.lean @@ -352,6 +352,16 @@ by ext n f R; refl end open_locale first_order + +instance constants_on_self_Structure : (constants_on M).Structure M := +constants_on.Structure id + +instance with_constants_self_Structure : L[[M]].Structure M := +language.sum_Structure _ _ M + +instance with_constants_self_expansion : (Lhom_with_constants L M).is_expansion_on M := +⟨λ _ _ _, rfl, λ _ _ _, rfl⟩ + variables (A : set M) instance with_constants_Structure : L[[A]].Structure M := From b337b921b0a15855924e88712d5fe51f95a08d1a Mon Sep 17 00:00:00 2001 From: Aaron Anderson Date: Wed, 4 May 2022 07:50:43 +0000 Subject: [PATCH 074/153] feat(model_theory/satisfiability): A union of a directed family of satisfiable theories is satisfiable (#13750) Proves `first_order.language.Theory.is_satisfiable_directed_union_iff` - the union of a directed family of theories is satisfiable if and only if all of the theories in the family are satisfiable. --- src/model_theory/satisfiability.lean | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/src/model_theory/satisfiability.lean b/src/model_theory/satisfiability.lean index 0f0b8745d9433..53f030d654c7a 100644 --- a/src/model_theory/satisfiability.lean +++ b/src/model_theory/satisfiability.lean @@ -87,6 +87,17 @@ theorem is_satisfiable_iff_is_finitely_satisfiable {T : L.Theory} : exact ⟨Model.of T M'⟩, end⟩ +theorem is_satisfiable_directed_union_iff {ι : Type*} [nonempty ι] + {T : ι → L.Theory} (h : directed (⊆) T) : + Theory.is_satisfiable (⋃ i, T i) ↔ ∀ i, (T i).is_satisfiable := +begin + refine ⟨λ h' i, h'.mono (set.subset_Union _ _), λ h', _⟩, + rw [is_satisfiable_iff_is_finitely_satisfiable, is_finitely_satisfiable], + intros T0 hT0, + obtain ⟨i, hi⟩ := h.exists_mem_subset_of_finset_subset_bUnion hT0, + exact (h' i).mono hi, +end + variable (T) /-- A theory models a (bounded) formula when any of its nonempty models realizes that formula on all From 9015d2adb19ee4166685d63ea54f5e97d026ac2e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 4 May 2022 07:50:44 +0000 Subject: [PATCH 075/153] refactor(set_theory/game/pgame): Redefine `subsequent` (#13752) We redefine `subsequent` as `trans_gen is_option`. This gives a much nicer induction principle than the previous one, and allows us to immediately prove well-foundedness. --- src/set_theory/game/pgame.lean | 57 ++++++++++++++++------------------ 1 file changed, 27 insertions(+), 30 deletions(-) diff --git a/src/set_theory/game/pgame.lean b/src/set_theory/game/pgame.lean index 87d9cdc483859..c0b48ba119363 100644 --- a/src/set_theory/game/pgame.lean +++ b/src/set_theory/game/pgame.lean @@ -5,6 +5,7 @@ Authors: Reid Barton, Mario Carneiro, Isabel Longbottom, Scott Morrison -/ import data.fin.basic import data.list.basic +import logic.relation /-! # Combinatorial (pre-)games. @@ -103,7 +104,7 @@ An interested reader may like to formalise some of the material from * [André Joyal, *Remarques sur la théorie des jeux à deux personnes*][joyal1997] -/ -open function +open function relation universes u @@ -204,43 +205,39 @@ theorem wf_is_option : well_founded is_option := { exact IHr j } end⟩ -/-- `subsequent p q` says that `p` can be obtained by playing - some nonempty sequence of moves from `q`. -/ -inductive subsequent : pgame → pgame → Prop -| left : Π (x : pgame) (i : x.left_moves), subsequent (x.move_left i) x -| right : Π (x : pgame) (j : x.right_moves), subsequent (x.move_right j) x -| trans : Π (x y z : pgame), subsequent x y → subsequent y z → subsequent x z - -theorem wf_subsequent : well_founded subsequent := -⟨λ x, begin - induction x with l r L R IHl IHr, - refine ⟨_, λ y h, _⟩, - generalize_hyp e : mk l r L R = x at h, - induction h with _ i _ j a b _ h1 h2 IH1 IH2; subst e, - { apply IHl }, - { apply IHr }, - { exact acc.inv (IH2 rfl) h1 } -end⟩ +/-- `subsequent x y` says that `x` can be obtained by playing some nonempty sequence of moves from +`y`. It is the transitive closure of `is_option`. -/ +def subsequent : pgame → pgame → Prop := +trans_gen is_option + +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 -instance : has_well_founded pgame := -{ r := subsequent, - wf := wf_subsequent } +instance : has_well_founded pgame := ⟨_, wf_subsequent⟩ -/-- A move by Left produces a subsequent game. (For use in pgame_wf_tac.) -/ -lemma subsequent.left_move {xl xr} {xL : xl → pgame} {xR : xr → pgame} {i : xl} : +lemma subsequent.move_left {x : pgame} (i : x.left_moves) : subsequent (x.move_left i) x := +trans_gen.single (is_option.move_left i) + +lemma subsequent.move_right {x : pgame} (j : x.right_moves) : subsequent (x.move_right j) x := +trans_gen.single (is_option.move_right j) + +lemma subsequent.mk_left {xl xr} (xL : xl → pgame) (xR : xr → pgame) (i : xl) : subsequent (xL i) (mk xl xr xL xR) := -subsequent.left (mk xl xr xL xR) i -/-- A move by Right produces a subsequent game. (For use in pgame_wf_tac.) -/ -lemma subsequent.right_move {xl xr} {xL : xl → pgame} {xR : xr → pgame} {j : xr} : +@subsequent.move_left (mk _ _ _ _) i + +lemma subsequent.mk_right {xl xr} (xL : xl → pgame) (xR : xr → pgame) (j : xr) : subsequent (xR j) (mk xl xr xL xR) := -subsequent.right (mk xl xr xL xR) j +@subsequent.move_right (mk _ _ _ _) j /-- A local tactic for proving well-foundedness of recursive definitions involving pregames. -/ meta def pgame_wf_tac := `[solve_by_elim - [psigma.lex.left, psigma.lex.right, - subsequent.left_move, subsequent.right_move, - subsequent.left, subsequent.right, subsequent.trans] + [psigma.lex.left, psigma.lex.right, subsequent.move_left, subsequent.move_right, + subsequent.mk_left, subsequent.mk_right, subsequent.trans] { max_depth := 6 }] /-- The pre-game `zero` is defined by `0 = { | }`. -/ From e3d38edbcbb728b03a9ffdc098b6e3b566d010b3 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 4 May 2022 07:50:45 +0000 Subject: [PATCH 076/153] feat(algebra/hom/non_unital_alg): some constructions for `prod` (#13785) --- src/algebra/hom/non_unital_alg.lean | 68 +++++++++++++++++++++++++++++ 1 file changed, 68 insertions(+) diff --git a/src/algebra/hom/non_unital_alg.lean b/src/algebra/hom/non_unital_alg.lean index f31095bbddca4..1c82398d8a8a3 100644 --- a/src/algebra/hom/non_unital_alg.lean +++ b/src/algebra/hom/non_unital_alg.lean @@ -189,8 +189,76 @@ def inverse (f : A →ₙₐ[R] B) (g : B → A) (inverse f g h₁ h₂ : B → A) = g := rfl +/-! ### Operations on the product type + +Note that much of this is copied from [`linear_algebra/prod`](../../linear_algebra/prod). -/ + +section prod + +variables (R A B) +/-- The first projection of a product is a non-unital alg_hom. -/ +@[simps] +def fst : A × B →ₙₐ[R] A := +{ to_fun := prod.fst, + map_zero' := rfl, map_add' := λ x y, rfl, map_smul' := λ x y, rfl, map_mul' := λ x y, rfl } + +/-- The second projection of a product is a non-unital alg_hom. -/ +@[simps] +def snd : A × B →ₙₐ[R] B := +{ to_fun := prod.snd, + map_zero' := rfl, map_add' := λ x y, rfl, map_smul' := λ x y, rfl, map_mul' := λ x y, rfl } + +variables {R A B} + +/-- The prod of two morphisms is a morphism. -/ +@[simps] def prod (f : A →ₙₐ[R] B) (g : A →ₙₐ[R] C) : (A →ₙₐ[R] B × C) := +{ to_fun := pi.prod f g, + map_zero' := by simp only [pi.prod, prod.zero_eq_mk, map_zero], + map_add' := λ x y, by simp only [pi.prod, prod.mk_add_mk, map_add], + map_mul' := λ x y, by simp only [pi.prod, prod.mk_mul_mk, map_mul], + map_smul' := λ c x, by simp only [pi.prod, prod.smul_mk, map_smul, ring_hom.id_apply] } + +lemma coe_prod (f : A →ₙₐ[R] B) (g : A →ₙₐ[R] C) : ⇑(f.prod g) = pi.prod f g := rfl + +@[simp] theorem fst_prod (f : A →ₙₐ[R] B) (g : A →ₙₐ[R] C) : + (fst R B C).comp (prod f g) = f := by ext; refl + +@[simp] theorem snd_prod (f : A →ₙₐ[R] B) (g : A →ₙₐ[R] C) : + (snd R B C).comp (prod f g) = g := by ext; refl + +@[simp] theorem prod_fst_snd : prod (fst R A B) (snd R A B) = 1 := +coe_injective pi.prod_fst_snd + +/-- Taking the product of two maps with the same domain is equivalent to taking the product of +their codomains. -/ +@[simps] def prod_equiv : ((A →ₙₐ[R] B) × (A →ₙₐ[R] C)) ≃ (A →ₙₐ[R] B × C) := +{ to_fun := λ f, f.1.prod f.2, + inv_fun := λ f, ((fst _ _ _).comp f, (snd _ _ _).comp f), + left_inv := λ f, by ext; refl, + right_inv := λ f, by ext; refl } + +variables (R A B) + +/-- The left injection into a product is a non-unital algebra homomorphism. -/ +def inl : A →ₙₐ[R] A × B := prod 1 0 + +/-- The right injection into a product is a non-unital algebra homomorphism. -/ +def inr : B →ₙₐ[R] A × B := prod 0 1 + +variables {R A B} + +@[simp] theorem coe_inl : (inl R A B : A → A × B) = λ x, (x, 0) := rfl +theorem inl_apply (x : A) : inl R A B x = (x, 0) := rfl + +@[simp] theorem coe_inr : (inr R A B : B → A × B) = prod.mk 0 := rfl +theorem inr_apply (x : B) : inr R A B x = (0, x) := rfl + +end prod + end non_unital_alg_hom +/-! ### Interaction with `alg_hom` -/ + namespace alg_hom variables {R A B} [comm_semiring R] [semiring A] [semiring B] [algebra R A] [algebra R B] From 85657f1b887a19de9aa4b6d2c326e15470b724fe Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 4 May 2022 07:50:46 +0000 Subject: [PATCH 077/153] feat(algebra/category/FinVect): has finite limits (#13793) Co-authored-by: Scott Morrison --- src/algebra/category/FinVect.lean | 18 +++-- src/algebra/category/FinVect/limits.lean | 70 +++++++++++++++++++ src/algebra/category/Module/products.lean | 56 +++++++++++++++ .../concrete_category/basic.lean | 9 +++ .../limits_of_products_and_equalizers.lean | 61 +++++++++++++--- src/linear_algebra/finite_dimensional.lean | 8 +++ 6 files changed, 209 insertions(+), 13 deletions(-) create mode 100644 src/algebra/category/FinVect/limits.lean create mode 100644 src/algebra/category/Module/products.lean diff --git a/src/algebra/category/FinVect.lean b/src/algebra/category/FinVect.lean index 7de3033e5d5f0..ce329a91a8816 100644 --- a/src/algebra/category/FinVect.lean +++ b/src/algebra/category/FinVect.lean @@ -11,8 +11,8 @@ import algebra.category.Module.monoidal /-! # The category of finite dimensional vector spaces -This introduces `FinVect K`, the category of finite dimensional vector spaces on a field `K`. -It is implemented as a full subcategory on a subtype of `Module K`. +This introduces `FinVect K`, the category of finite dimensional vector spaces over a field `K`. +It is implemented as a full subcategory on a subtype of `Module K`. We first create the instance as a category, then as a monoidal category and then as a rigid monoidal category. @@ -31,12 +31,12 @@ universes u variables (K : Type u) [field K] /-- Define `FinVect` as the subtype of `Module.{u} K` of finite dimensional vector spaces. -/ -@[derive [category, λ α, has_coe_to_sort α (Sort*)]] +@[derive [large_category, λ α, has_coe_to_sort α (Sort*), concrete_category]] def FinVect := { V : Module.{u} K // finite_dimensional K V } namespace FinVect -instance finite_dimensional (V : FinVect K): finite_dimensional K V := V.prop +instance finite_dimensional (V : FinVect K) : finite_dimensional K V := V.prop instance : inhabited (FinVect K) := ⟨⟨Module.of K K, finite_dimensional.finite_dimensional_self K⟩⟩ @@ -45,6 +45,16 @@ instance : has_coe (FinVect.{u} K) (Module.{u} K) := { coe := λ V, V.1, } protected lemma coe_comp {U V W : FinVect K} (f : U ⟶ V) (g : V ⟶ W) : ((f ≫ g) : U → W) = (g : V → W) ∘ (f : U → V) := rfl +/-- Lift an unbundled vector space to `FinVect K`. -/ +def of (V : Type u) [add_comm_group V] [module K V] [finite_dimensional K V] : FinVect K := +⟨Module.of K V, by { change finite_dimensional K V, apply_instance }⟩ + +instance : has_forget₂ (FinVect.{u} K) (Module.{u} K) := +by { dsimp [FinVect], apply_instance, } + +instance : full (forget₂ (FinVect K) (Module.{u} K)) := +{ preimage := λ X Y f, f, } + instance monoidal_category : monoidal_category (FinVect K) := monoidal_category.full_monoidal_subcategory (λ V, finite_dimensional K V) diff --git a/src/algebra/category/FinVect/limits.lean b/src/algebra/category/FinVect/limits.lean new file mode 100644 index 0000000000000..fff72cf305aa6 --- /dev/null +++ b/src/algebra/category/FinVect/limits.lean @@ -0,0 +1,70 @@ +/- +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.category.FinVect +import algebra.category.Module.limits +import algebra.category.Module.products +import algebra.category.Module.epi_mono +import category_theory.limits.creates +import category_theory.limits.shapes.finite_limits +import category_theory.limits.constructions.limits_of_products_and_equalizers + +/-! +# `forget₂ (FinVect K) (Module K)` creates all finite limits. + +And hence `FinVect K` has all finite limits. +-/ + +noncomputable theory +universes v u + +open category_theory +open category_theory.limits + +namespace FinVect + +variables {J : Type v} [small_category J] [fin_category J] +variables {k : Type v} [field k] + +instance {J : Type v} [fintype J] (Z : J → Module.{v} k) [∀ j, finite_dimensional k (Z j)] : + finite_dimensional k (∏ λ j, Z j : Module.{v} k) := +begin + haveI : finite_dimensional k (Module.of k (Π j, Z j)), { dsimp, apply_instance, }, + exact finite_dimensional.of_injective + (Module.pi_iso_pi _).hom + ((Module.mono_iff_injective _).1 (by apply_instance)), +end + +/-- Finite limits of finite finite dimensional vectors spaces are finite dimensional, +because we can realise them as subobjects of a finite product. -/ +instance (F : J ⥤ FinVect k) : + finite_dimensional k (limit (F ⋙ forget₂ (FinVect k) (Module.{v} k)) : Module.{v} k) := +begin + haveI : ∀ j, finite_dimensional k ((F ⋙ forget₂ (FinVect k) (Module.{v} k)).obj j), + { intro j, change finite_dimensional k (F.obj j), apply_instance, }, + exact finite_dimensional.of_injective + (limit_subobject_product (F ⋙ forget₂ (FinVect k) (Module.{v} k))) + ((Module.mono_iff_injective _).1 (by apply_instance)), +end + +/-- The forgetful functor from `FinVect k` to `Module k` creates all finite limits. -/ +def forget₂_creates_limit (F : J ⥤ FinVect k) : + creates_limit F (forget₂ (FinVect k) (Module.{v} k)) := +creates_limit_of_fully_faithful_of_iso + ⟨(limit (F ⋙ forget₂ (FinVect k) (Module.{v} k)) : Module.{v} k), by apply_instance⟩ + (iso.refl _) + +instance : creates_limits_of_shape J (forget₂ (FinVect k) (Module.{v} k)) := +{ creates_limit := λ F, forget₂_creates_limit F, } + +instance : has_finite_limits (FinVect k) := +{ out := λ J _ _, by exactI + has_limits_of_shape_of_has_limits_of_shape_creates_limits_of_shape + (forget₂ (FinVect k) (Module.{v} k)), } + +instance : preserves_finite_limits (forget₂ (FinVect k) (Module.{v} k)) := +{ preserves_finite_limits := λ J _ _, by exactI infer_instance, } + +end FinVect diff --git a/src/algebra/category/Module/products.lean b/src/algebra/category/Module/products.lean new file mode 100644 index 0000000000000..18c1c7fd3194f --- /dev/null +++ b/src/algebra/category/Module/products.lean @@ -0,0 +1,56 @@ +/- +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.category.Module.epi_mono +import linear_algebra.pi + +/-! +# The concrete products in the category of modules are products in the categorical sense. +-/ + +open category_theory +open category_theory.limits + +universes u v + +namespace Module +variables {R : Type u} [ring R] + +variables {ι : Type v} (Z : ι → Module.{v} R) + +/-- The product cone induced by the concrete product. -/ +def product_cone : fan Z := +fan.mk (Module.of R (Π i : ι, Z i)) (λ i, (linear_map.proj i : (Π i : ι, Z i) →ₗ[R] Z i)) + +/-- The concrete product cone is limiting. -/ +def product_cone_is_limit : is_limit (product_cone Z) := +{ lift := λ s, (linear_map.pi s.π.app : s.X →ₗ[R] (Π i : ι, Z i)), + fac' := by tidy, + uniq' := λ s m w, by { ext x i, exact linear_map.congr_fun (w i) x, }, } + +-- While we could use this to construct a `has_products (Module R)` instance, +-- we already have `has_limits (Module R)` in `algebra.category.Module.limits`. + +variables [has_products (Module.{v} R)] + +/-- +The categorical product of a family of objects in `Module` +agrees with the usual module-theoretical product. +-/ +noncomputable def pi_iso_pi : + ∏ Z ≅ Module.of R (Π i, Z i) := +limit.iso_limit_cone ⟨_, product_cone_is_limit Z⟩ + +-- We now show this isomorphism commutes with the inclusion of the kernel into the source. + +@[simp, elementwise] lemma pi_iso_pi_inv_kernel_ι (i : ι) : + (pi_iso_pi Z).inv ≫ pi.π Z i = (linear_map.proj i : (Π i : ι, Z i) →ₗ[R] Z i) := +limit.iso_limit_cone_inv_π _ _ + +@[simp, elementwise] lemma pi_iso_pi_hom_ker_subtype (i : ι) : + (pi_iso_pi Z).hom ≫ (linear_map.proj i : (Π i : ι, Z i) →ₗ[R] Z i) = pi.π Z i := +is_limit.cone_point_unique_up_to_iso_inv_comp _ (limit.is_limit _) _ + +end Module diff --git a/src/category_theory/concrete_category/basic.lean b/src/category_theory/concrete_category/basic.lean index 00524c6e2b8f4..a565d66e9e774 100644 --- a/src/category_theory/concrete_category/basic.lean +++ b/src/category_theory/concrete_category/basic.lean @@ -177,6 +177,15 @@ instance induced_category.has_forget₂ {C : Type v} {D : Type v'} [category D] { forget₂ := induced_functor f, forget_comp := rfl } +instance full_subcategory.concrete_category {C : Type v} [category C] [concrete_category C] + (Z : C → Prop) : concrete_category {X : C // Z X} := +{ forget := full_subcategory_inclusion Z ⋙ forget C } + +instance full_subcategory.has_forget₂ {C : Type v} [category C] [concrete_category C] + (Z : C → Prop) : has_forget₂ {X : C // Z X} C := +{ forget₂ := full_subcategory_inclusion Z, + forget_comp := rfl } + /-- In order to construct a “partially forgetting” functor, we do not need to verify functor laws; it suffices to ensure that compositions agree with `forget₂ C D ⋙ forget D = forget C`. diff --git a/src/category_theory/limits/constructions/limits_of_products_and_equalizers.lean b/src/category_theory/limits/constructions/limits_of_products_and_equalizers.lean index eebdabe7770a3..910463e7a3a9f 100644 --- a/src/category_theory/limits/constructions/limits_of_products_and_equalizers.lean +++ b/src/category_theory/limits/constructions/limits_of_products_and_equalizers.lean @@ -82,15 +82,15 @@ end has_limit_of_has_products_of_has_equalizers open has_limit_of_has_products_of_has_equalizers /-- -Given the existence of the appropriate (possibly finite) products and equalizers, we know a limit of -`F` exists. +Given the existence of the appropriate (possibly finite) products and equalizers, +we can construct a limit cone for `F`. (This assumes the existence of all equalizers, which is technically stronger than needed.) -/ -lemma has_limit_of_equalizer_and_product (F : J ⥤ C) +noncomputable +def limit_cone_of_equalizer_and_product (F : J ⥤ C) [has_limit (discrete.functor F.obj)] [has_limit (discrete.functor (λ f : (Σ p : J × J, p.1 ⟶ p.2), F.obj f.1.2))] - [has_equalizers C] : has_limit F := -has_limit.mk + [has_equalizers C] : limit_cone F := { cone := _, is_limit := build_is_limit @@ -102,6 +102,27 @@ has_limit.mk (limit.is_limit _) (limit.is_limit _) } +/-- +Given the existence of the appropriate (possibly finite) products and equalizers, we know a limit of +`F` exists. +(This assumes the existence of all equalizers, which is technically stronger than needed.) +-/ +lemma has_limit_of_equalizer_and_product (F : J ⥤ C) + [has_limit (discrete.functor F.obj)] + [has_limit (discrete.functor (λ f : (Σ p : J × J, p.1 ⟶ p.2), F.obj f.1.2))] + [has_equalizers C] : has_limit F := +has_limit.mk (limit_cone_of_equalizer_and_product F) + +/-- A limit can be realised as a subobject of a product. -/ +noncomputable +def limit_subobject_product [has_limits C] (F : J ⥤ C) : + limit F ⟶ ∏ (λ j, F.obj j) := +(limit.iso_limit_cone (limit_cone_of_equalizer_and_product F)).hom ≫ equalizer.ι _ _ + +instance limit_subobject_product_mono [has_limits C] (F : J ⥤ C) : + mono (limit_subobject_product F) := +mono_comp _ _ + /-- Any category with products and equalizers has all limits. @@ -244,14 +265,14 @@ open has_colimit_of_has_coproducts_of_has_coequalizers /-- Given the existence of the appropriate (possibly finite) coproducts and coequalizers, -we know a colimit of `F` exists. +we can construct a colimit cocone for `F`. (This assumes the existence of all coequalizers, which is technically stronger than needed.) -/ -lemma has_colimit_of_coequalizer_and_coproduct (F : J ⥤ C) +noncomputable +def colimit_cocone_of_coequalizer_and_coproduct (F : J ⥤ C) [has_colimit (discrete.functor F.obj)] [has_colimit (discrete.functor (λ f : (Σ p : J × J, p.1 ⟶ p.2), F.obj f.1.1))] - [has_coequalizers C] : has_colimit F := -has_colimit.mk + [has_coequalizers C] : colimit_cocone F := { cocone := _, is_colimit := build_is_colimit @@ -263,6 +284,28 @@ has_colimit.mk (colimit.is_colimit _) (colimit.is_colimit _) } + +/-- +Given the existence of the appropriate (possibly finite) coproducts and coequalizers, +we know a colimit of `F` exists. +(This assumes the existence of all coequalizers, which is technically stronger than needed.) +-/ +lemma has_colimit_of_coequalizer_and_coproduct (F : J ⥤ C) + [has_colimit (discrete.functor F.obj)] + [has_colimit (discrete.functor (λ f : (Σ p : J × J, p.1 ⟶ p.2), F.obj f.1.1))] + [has_coequalizers C] : has_colimit F := +has_colimit.mk (colimit_cocone_of_coequalizer_and_coproduct F) + +/-- A colimit can be realised as a quotient of a coproduct. -/ +noncomputable +def colimit_quotient_coproduct [has_colimits C] (F : J ⥤ C) : + ∐ (λ j, F.obj j) ⟶ colimit F := +coequalizer.π _ _ ≫ (colimit.iso_colimit_cocone (colimit_cocone_of_coequalizer_and_coproduct F)).inv + +instance colimit_quotient_coproduct_epi [has_colimits C] (F : J ⥤ C) : + epi (colimit_quotient_coproduct F) := +epi_comp _ _ + /-- Any category with coproducts and coequalizers has all colimits. diff --git a/src/linear_algebra/finite_dimensional.lean b/src/linear_algebra/finite_dimensional.lean index bfb42092bc22c..3fda0e8713624 100644 --- a/src/linear_algebra/finite_dimensional.lean +++ b/src/linear_algebra/finite_dimensional.lean @@ -108,6 +108,14 @@ variables (K V) instance finite_dimensional_pi {ι} [fintype ι] : finite_dimensional K (ι → K) := iff_fg.1 is_noetherian_pi +instance finite_dimensional_pi' {ι} [fintype ι] (M : ι → Type*) + [∀ i, add_comm_group (M i)] [∀ i, module K (M i)] [I : ∀ i, finite_dimensional K (M i)] : + finite_dimensional K (Π i, M i) := +begin + haveI : ∀ i : ι, is_noetherian K (M i) := λ i, iff_fg.2 (I i), + exact iff_fg.1 is_noetherian_pi +end + /-- A finite dimensional vector space over a finite field is finite -/ noncomputable def fintype_of_fintype [fintype K] [finite_dimensional K V] : fintype V := module.fintype_of_fintype (@finset_basis K V _ _ _ (iff_fg.2 infer_instance)) From ba4bf54514b4ba6ed84cbc226c5d5e45762cea0c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 4 May 2022 07:50:47 +0000 Subject: [PATCH 078/153] feat(set_theory/game/pgame): Add more congr lemmas (#13808) Co-authored-by: Johan Commelin --- src/set_theory/game/pgame.lean | 33 ++++++++++++++++++++++----------- 1 file changed, 22 insertions(+), 11 deletions(-) diff --git a/src/set_theory/game/pgame.lean b/src/set_theory/game/pgame.lean index c0b48ba119363..e423d78dd20dc 100644 --- a/src/set_theory/game/pgame.lean +++ b/src/set_theory/game/pgame.lean @@ -426,15 +426,15 @@ end @[simp] theorem not_le {x y : pgame} : ¬ x ≤ y ↔ y < x := not_le_lt.1 @[simp] theorem not_lt {x y : pgame} : ¬ x < y ↔ y ≤ x := not_le_lt.2 -@[refl] protected theorem le_refl : ∀ x : pgame, x ≤ x +@[refl] protected theorem le_rfl : ∀ {x : pgame}, x ≤ x | ⟨l, r, L, R⟩ := by rw mk_le_mk; exact -⟨λ i, lt_mk_of_le (le_refl _), λ i, mk_lt_of_le (le_refl _)⟩ +⟨λ i, lt_mk_of_le le_rfl, λ i, mk_lt_of_le le_rfl⟩ -protected theorem le_rfl {x : pgame} : x ≤ x := -pgame.le_refl x +protected theorem le_refl (x : pgame) : x ≤ x := +pgame.le_rfl protected theorem lt_irrefl (x : pgame) : ¬ x < x := -not_lt.2 (pgame.le_refl _) +not_lt.2 pgame.le_rfl protected theorem ne_of_lt : ∀ {x y : pgame}, x < y → x ≠ y | x _ h rfl := pgame.lt_irrefl x h @@ -505,7 +505,8 @@ def equiv (x y : pgame) : Prop := x ≤ y ∧ y ≤ x local infix ` ≈ ` := pgame.equiv -@[refl, simp] theorem equiv_refl (x) : x ≈ x := ⟨pgame.le_refl _, pgame.le_refl _⟩ +@[refl, simp] theorem equiv_rfl {x} : x ≈ x := ⟨pgame.le_rfl, pgame.le_rfl⟩ +theorem equiv_refl (x) : x ≈ x := equiv_rfl @[symm] theorem equiv_symm {x y} : x ≈ y → y ≈ x | ⟨xy, yx⟩ := ⟨yx, xy⟩ @[trans] theorem equiv_trans {x y z} : x ≈ y → y ≈ z → x ≈ z | ⟨xy, yx⟩ ⟨yz, zy⟩ := ⟨le_trans xy yz, le_trans zy yx⟩ @@ -971,7 +972,7 @@ using_well_founded { dec_tac := pgame_wf_tac } theorem neg_add_le {x y : pgame} : -(x + y) ≤ -x + -y := (neg_add_relabelling x y).le -/-- `x+y` has exactly the same moves as `y+x`. -/ +/-- `x + y` has exactly the same moves as `y + x`. -/ def add_comm_relabelling : Π (x y : pgame.{u}), relabelling (x + y) (y + x) | (mk xl xr xL xR) (mk yl yr yL yR) := begin @@ -1072,14 +1073,24 @@ instance covariant_class_add_le : covariant_class pgame pgame (+) (≤) := ... ≤ x + z : add_comm_le⟩ theorem add_congr {w x y z : pgame} (h₁ : w ≈ x) (h₂ : y ≈ z) : w + y ≈ x + z := -⟨calc w + y ≤ w + z : add_le_add_left h₂.1 _ - ... ≤ x + z : add_le_add_right h₁.1 _, - calc x + z ≤ x + y : add_le_add_left h₂.2 _ - ... ≤ w + y : add_le_add_right h₁.2 _⟩ +⟨le_trans (add_le_add_left h₂.1 w) (add_le_add_right h₁.1 z), + le_trans (add_le_add_left h₂.2 x) (add_le_add_right h₁.2 y)⟩ + +theorem add_congr_left {x y z : pgame} (h : x ≈ y) : x + z ≈ y + z := +add_congr h equiv_rfl + +theorem add_congr_right {x y z : pgame} : y ≈ z → x + y ≈ x + z := +add_congr equiv_rfl theorem sub_congr {w x y z : pgame} (h₁ : w ≈ x) (h₂ : y ≈ z) : w - y ≈ x - z := add_congr h₁ (neg_congr h₂) +theorem sub_congr_left {x y z : pgame} (h : x ≈ y) : x - z ≈ y - z := +sub_congr h equiv_rfl + +theorem sub_congr_right {x y z : pgame} : y ≈ z → x - y ≈ x - z := +sub_congr equiv_rfl + theorem add_left_neg_le_zero : ∀ (x : pgame), -x + x ≤ 0 | ⟨xl, xr, xL, xR⟩ := begin From 6e0033066b049b9689ef6d3aa7e2970e8e175385 Mon Sep 17 00:00:00 2001 From: Bhavik Mehta Date: Wed, 4 May 2022 07:50:48 +0000 Subject: [PATCH 079/153] feat(algebra/squarefree): relate squarefree on naturals to factorization (#13816) Also moves `nat.two_le_iff` higher up the hierarchy since it's an elementary lemma and give it a more appropriate type. The lemma `squarefree_iff_prime_sq_not_dvd` has been deleted because it's a duplicate of a lemma which is already earlier in the same file. Co-authored-by: Bhavik Mehta --- src/algebra/squarefree.lean | 74 ++++++++++++++++++++++++++++++--- src/data/nat/basic.lean | 5 +++ src/data/nat/factorization.lean | 7 ++++ src/data/nat/prime.lean | 6 --- 4 files changed, 81 insertions(+), 11 deletions(-) diff --git a/src/algebra/squarefree.lean b/src/algebra/squarefree.lean index 12c7f5a4f031b..21ef0efe6a34f 100644 --- a/src/algebra/squarefree.lean +++ b/src/algebra/squarefree.lean @@ -227,6 +227,75 @@ end theorem squarefree_iff_prime_squarefree {n : ℕ} : squarefree n ↔ ∀ x, prime x → ¬ x * x ∣ n := squarefree_iff_irreducible_sq_not_dvd_of_exists_irreducible ⟨_, prime_two⟩ +lemma squarefree.factorization_le_one {n : ℕ} (p : ℕ) (hn : squarefree n) : + n.factorization p ≤ 1 := +begin + rcases eq_or_ne n 0 with rfl | hn', + { simp }, + rw [multiplicity.squarefree_iff_multiplicity_le_one] at hn, + by_cases hp : p.prime, + { have := hn p, + simp only [multiplicity_eq_factorization hp hn', nat.is_unit_iff, hp.ne_one, or_false] at this, + exact_mod_cast this }, + { rw factorization_eq_zero_of_non_prime _ hp, + exact zero_le_one } +end + +lemma squarefree_of_factorization_le_one {n : ℕ} (hn : n ≠ 0) (hn' : ∀ p, n.factorization p ≤ 1) : + squarefree n := +begin + rw [squarefree_iff_nodup_factors hn, list.nodup_iff_count_le_one], + intro a, + rw factors_count_eq, + apply hn', +end + +lemma squarefree_iff_factorization_le_one {n : ℕ} (hn : n ≠ 0) : + squarefree n ↔ ∀ p, n.factorization p ≤ 1 := +⟨λ p hn, squarefree.factorization_le_one hn p, squarefree_of_factorization_le_one hn⟩ + +lemma squarefree.ext_iff {n m : ℕ} (hn : squarefree n) (hm : squarefree m) : + n = m ↔ ∀ p, prime p → (p ∣ n ↔ p ∣ m) := +begin + refine ⟨by { rintro rfl, simp }, λ h, eq_of_factorization_eq hn.ne_zero hm.ne_zero (λ p, _)⟩, + by_cases hp : p.prime, + { have h₁ := h _ hp, + rw [←not_iff_not, hp.dvd_iff_one_le_factorization hn.ne_zero, not_le, lt_one_iff, + hp.dvd_iff_one_le_factorization hm.ne_zero, not_le, lt_one_iff] at h₁, + have h₂ := squarefree.factorization_le_one p hn, + have h₃ := squarefree.factorization_le_one p hm, + rw [nat.le_add_one_iff, le_zero_iff] at h₂ h₃, + cases h₂, + { rwa [h₂, eq_comm, ←h₁] }, + { rw [h₂, h₃.resolve_left], + rw [←h₁, h₂], + simp only [nat.one_ne_zero, not_false_iff] } }, + rw [factorization_eq_zero_of_non_prime _ hp, factorization_eq_zero_of_non_prime _ hp], +end + +lemma squarefree_pow_iff {n k : ℕ} (hn : n ≠ 1) (hk : k ≠ 0) : + squarefree (n ^ k) ↔ squarefree n ∧ k = 1 := +begin + refine ⟨λ h, _, by { rintro ⟨hn, rfl⟩, simpa }⟩, + rcases eq_or_ne n 0 with rfl | hn₀, + { simpa [zero_pow hk.bot_lt] using h }, + refine ⟨squarefree_of_dvd_of_squarefree (dvd_pow_self _ hk) h, by_contradiction $ λ h₁, _⟩, + have : 2 ≤ k := k.two_le_iff.mpr ⟨hk, h₁⟩, + apply hn (nat.is_unit_iff.1 (h _ _)), + rw ←sq, + exact pow_dvd_pow _ this +end + +lemma squarefree_and_prime_pow_iff_prime {n : ℕ} : + squarefree n ∧ is_prime_pow n ↔ prime n := +begin + refine iff.symm ⟨λ hn, ⟨hn.squarefree, hn.is_prime_pow⟩, _⟩, + rw is_prime_pow_nat_iff, + rintro ⟨h, p, k, hp, hk, rfl⟩, + rw squarefree_pow_iff hp.ne_one hk.ne' at h, + rwa [h.2, pow_one], +end + /-- Assuming that `n` has no factors less than `k`, returns the smallest prime `p` such that `p^2 ∣ n`. -/ def min_sq_fac_aux : ℕ → ℕ → option ℕ @@ -459,11 +528,6 @@ begin exact ⟨a, b, h₁, h₂⟩ }, end -lemma squarefree_iff_prime_sq_not_dvd (n : ℕ) : - squarefree n ↔ ∀ x : ℕ, x.prime → ¬ x * x ∣ n := -squarefree_iff_irreducible_sq_not_dvd_of_exists_irreducible - ⟨2, (irreducible_iff_nat_prime _).2 prime_two⟩ - /-- `squarefree` is multiplicative. Note that the → direction does not require `hmn` and generalizes to arbitrary commutative monoids. See `squarefree.of_mul_left` and `squarefree.of_mul_right` above for auxiliary lemmas. -/ diff --git a/src/data/nat/basic.lean b/src/data/nat/basic.lean index 6903ff1ac2f78..739da4899d94f 100644 --- a/src/data/nat/basic.lean +++ b/src/data/nat/basic.lean @@ -315,6 +315,11 @@ succ_ne_succ.mpr n.succ_ne_zero @[simp] lemma one_lt_succ_succ (n : ℕ) : 1 < n.succ.succ := succ_lt_succ $ succ_pos n +lemma two_le_iff : ∀ n, 2 ≤ n ↔ n ≠ 0 ∧ n ≠ 1 +| 0 := by simp +| 1 := by simp +| (n+2) := by simp + theorem succ_le_succ_iff {m n : ℕ} : succ m ≤ succ n ↔ m ≤ n := ⟨le_of_succ_le_succ, succ_le_succ⟩ diff --git a/src/data/nat/factorization.lean b/src/data/nat/factorization.lean index b61a536d91e26..89dc6eea89007 100644 --- a/src/data/nat/factorization.lean +++ b/src/data/nat/factorization.lean @@ -93,6 +93,9 @@ le_of_mem_factors (factor_iff_mem_factorization.mp h) lemma factorization_eq_zero_of_non_prime (n : ℕ) {p : ℕ} (hp : ¬p.prime) : n.factorization p = 0 := not_mem_support_iff.1 (mt prime_of_mem_factorization hp) +lemma dvd_of_factorization_pos {n p : ℕ} (hn : n.factorization p ≠ 0) : p ∣ n := +dvd_of_mem_factors (factor_iff_mem_factorization.1 (mem_support_iff.2 hn)) + lemma prime.factorization_pos_of_dvd {n p : ℕ} (hp : p.prime) (hn : n ≠ 0) (h : p ∣ n) : 0 < n.factorization p := by rwa [←factors_count_eq, count_pos, mem_factors_iff_dvd hn hp] @@ -255,6 +258,10 @@ lemma prime.pow_dvd_iff_dvd_pow_factorization {p k n : ℕ} (pp : prime p) (hn : p ^ k ∣ n ↔ p ^ k ∣ p ^ n.factorization p := by rw [pow_dvd_pow_iff_le_right pp.one_lt, pp.pow_dvd_iff_le_factorization hn] +lemma prime.dvd_iff_one_le_factorization {p n : ℕ} (pp : prime p) (hn : n ≠ 0) : + p ∣ n ↔ 1 ≤ n.factorization p := +iff.trans (by simp) (pp.pow_dvd_iff_le_factorization hn) + lemma exists_factorization_lt_of_lt {a b : ℕ} (ha : a ≠ 0) (hab : a < b) : ∃ p : ℕ, a.factorization p < b.factorization p := begin diff --git a/src/data/nat/prime.lean b/src/data/nat/prime.lean index 6800c71ac9186..ea7845c9b5f43 100644 --- a/src/data/nat/prime.lean +++ b/src/data/nat/prime.lean @@ -64,12 +64,6 @@ instance prime.one_lt' (p : ℕ) [hp : _root_.fact p.prime] : _root_.fact (1 < p lemma prime.ne_one {p : ℕ} (hp : p.prime) : p ≠ 1 := hp.one_lt.ne' -lemma two_le_iff (n : ℕ) : 2 ≤ n ↔ n ≠ 0 ∧ ¬is_unit n := -begin - rw nat.is_unit_iff, - rcases n with _|_|m; norm_num [one_lt_succ_succ, succ_le_iff] -end - lemma prime.eq_one_or_self_of_dvd {p : ℕ} (pp : p.prime) (m : ℕ) (hm : m ∣ p) : m = 1 ∨ m = p := begin obtain ⟨n, hn⟩ := hm, From 35c8980d95a735a176fff87a5619f3fd1a97ef7d Mon Sep 17 00:00:00 2001 From: sgouezel Date: Wed, 4 May 2022 07:50:49 +0000 Subject: [PATCH 080/153] feat(analysis/asymptotics/specific_asymptotics): Cesaro averaging preserves convergence (#13825) --- .../asymptotics/specific_asymptotics.lean | 78 +++++++++++++++++++ 1 file changed, 78 insertions(+) diff --git a/src/analysis/asymptotics/specific_asymptotics.lean b/src/analysis/asymptotics/specific_asymptotics.lean index 3e7704fbccb4c..3e1d8c52e0e4c 100644 --- a/src/analysis/asymptotics/specific_asymptotics.lean +++ b/src/analysis/asymptotics/specific_asymptotics.lean @@ -98,3 +98,81 @@ begin end end normed_linear_ordered_field + +section real + +open_locale big_operators +open finset + +lemma asymptotics.is_o.sum_range {α : Type*} [normed_group α] + {f : ℕ → α} {g : ℕ → ℝ} (h : is_o f g at_top) + (hg : 0 ≤ g) (h'g : tendsto (λ n, ∑ i in range n, g i) at_top at_top) : + is_o (λ n, ∑ i in range n, f i) (λ n, ∑ i in range n, g i) at_top := +begin + have A : ∀ i, ∥g i∥ = g i := λ i, real.norm_of_nonneg (hg i), + have B : ∀ n, ∥∑ i in range n, g i∥ = ∑ i in range n, g i, + from λ n, by rwa [real.norm_eq_abs, abs_sum_of_nonneg'], + apply is_o_iff.2 (λ ε εpos, _), + obtain ⟨N, hN⟩ : ∃ (N : ℕ), ∀ (b : ℕ), N ≤ b → ∥f b∥ ≤ ε / 2 * g b, + by simpa only [A, eventually_at_top] using is_o_iff.mp h (half_pos εpos), + have : is_o (λ (n : ℕ), ∑ i in range N, f i) (λ (n : ℕ), ∑ i in range n, g i) at_top, + { apply is_o_const_left.2, + exact or.inr (h'g.congr (λ n, (B n).symm)) }, + filter_upwards [is_o_iff.1 this (half_pos εpos), Ici_mem_at_top N] with n hn Nn, + calc ∥∑ i in range n, f i∥ + = ∥∑ i in range N, f i + ∑ i in Ico N n, f i∥ : + by rw sum_range_add_sum_Ico _ Nn + ... ≤ ∥∑ i in range N, f i∥ + ∥∑ i in Ico N n, f i∥ : + norm_add_le _ _ + ... ≤ ∥∑ i in range N, f i∥ + ∑ i in Ico N n, (ε / 2) * g i : + add_le_add le_rfl (norm_sum_le_of_le _ (λ i hi, hN _ (mem_Ico.1 hi).1)) + ... ≤ ∥∑ i in range N, f i∥ + ∑ i in range n, (ε / 2) * g i : + begin + refine add_le_add le_rfl _, + apply sum_le_sum_of_subset_of_nonneg, + { rw range_eq_Ico, + exact Ico_subset_Ico (zero_le _) le_rfl }, + { assume i hi hident, + exact mul_nonneg (half_pos εpos).le (hg i) } + end + ... ≤ (ε / 2) * ∥∑ i in range n, g i∥ + (ε / 2) * (∑ i in range n, g i) : + begin + rw ← mul_sum, + exact add_le_add hn (mul_le_mul_of_nonneg_left le_rfl (half_pos εpos).le), + end + ... = ε * ∥(∑ i in range n, g i)∥ : by { simp [B], ring } +end + +lemma asymptotics.is_o_sum_range_of_tendsto_zero {α : Type*} [normed_group α] + {f : ℕ → α} (h : tendsto f at_top (𝓝 0)) : + is_o (λ n, ∑ i in range n, f i) (λ n, (n : ℝ)) at_top := +begin + have := ((is_o_one_iff ℝ).2 h).sum_range (λ i, zero_le_one), + simp only [sum_const, card_range, nat.smul_one_eq_coe] at this, + exact this tendsto_coe_nat_at_top_at_top +end + +/-- The Cesaro average of a converging sequence converges to the same limit. -/ +lemma filter.tendsto.cesaro_smul {E : Type*} [normed_group E] [normed_space ℝ E] + {u : ℕ → E} {l : E} (h : tendsto u at_top (𝓝 l)) : + tendsto (λ (n : ℕ), (n ⁻¹ : ℝ) • (∑ i in range n, u i)) at_top (𝓝 l) := +begin + rw [← tendsto_sub_nhds_zero_iff, ← is_o_one_iff ℝ], + have := asymptotics.is_o_sum_range_of_tendsto_zero (tendsto_sub_nhds_zero_iff.2 h), + apply ((is_O_refl (λ (n : ℕ), (n : ℝ) ⁻¹) at_top).smul_is_o this).congr' _ _, + { filter_upwards [Ici_mem_at_top 1] with n npos, + have nposℝ : (0 : ℝ) < n := nat.cast_pos.2 npos, + simp only [smul_sub, sum_sub_distrib, sum_const, card_range, sub_right_inj], + rw [nsmul_eq_smul_cast ℝ, smul_smul, inv_mul_cancel nposℝ.ne', one_smul] }, + { filter_upwards [Ici_mem_at_top 1] with n npos, + have nposℝ : (0 : ℝ) < n := nat.cast_pos.2 npos, + rw [algebra.id.smul_eq_mul, inv_mul_cancel nposℝ.ne'] } +end + +/-- The Cesaro average of a converging sequence converges to the same limit. -/ +lemma filter.tendsto.cesaro + {u : ℕ → ℝ} {l : ℝ} (h : tendsto u at_top (𝓝 l)) : + tendsto (λ (n : ℕ), (n ⁻¹ : ℝ) * (∑ i in range n, u i)) at_top (𝓝 l) := +h.cesaro_smul + +end real From 517aa8b6f54a2730709a3f294b4306d3f0de4ee2 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 4 May 2022 07:50:50 +0000 Subject: [PATCH 081/153] feat(topology/algebra/star): continuity of `star` (#13855) This adds the obvious instances for `pi`, `prod`, `units`, `mul_opposite`, `real`, `complex`, `is_R_or_C`, and `matrix`. We already had a `continuous_star` lemma, but it had stronger typeclass assumptions. This resolves multiple TODO comments. Co-authored-by: Johan Commelin --- src/analysis/complex/basic.lean | 4 +- src/analysis/normed_space/star/basic.lean | 33 +------ src/data/complex/is_R_or_C.lean | 5 +- src/topology/algebra/star.lean | 86 +++++++++++++++++++ .../continuous_function/zero_at_infty.lean | 42 +++++---- src/topology/instances/matrix.lean | 39 +++++++-- src/topology/instances/real.lean | 3 + 7 files changed, 157 insertions(+), 55 deletions(-) create mode 100644 src/topology/algebra/star.lean diff --git a/src/analysis/complex/basic.lean b/src/analysis/complex/basic.lean index 378408840be86..ee6c02c4604af 100644 --- a/src/analysis/complex/basic.lean +++ b/src/analysis/complex/basic.lean @@ -197,7 +197,9 @@ det_conj_ae @[simp] lemma linear_equiv_det_conj_lie : conj_lie.to_linear_equiv.det = -1 := linear_equiv_det_conj_ae -@[continuity] lemma continuous_conj : continuous (conj : ℂ → ℂ) := conj_lie.continuous +instance : has_continuous_star ℂ := ⟨conj_lie.continuous⟩ + +@[continuity] lemma continuous_conj : continuous (conj : ℂ → ℂ) := continuous_star /-- Continuous linear equiv version of the conj function, from `ℂ` to `ℂ`. -/ def conj_cle : ℂ ≃L[ℝ] ℂ := conj_lie diff --git a/src/analysis/normed_space/star/basic.lean b/src/analysis/normed_space/star/basic.lean index 84eaf1fd7220c..864193d1e9b1d 100644 --- a/src/analysis/normed_space/star/basic.lean +++ b/src/analysis/normed_space/star/basic.lean @@ -57,36 +57,9 @@ def star_normed_group_hom : normed_group_hom E E := lemma star_isometry : isometry (star : E → E) := star_add_equiv.to_add_monoid_hom.isometry_of_norm norm_star -lemma continuous_star : continuous (star : E → E) := star_isometry.continuous - -lemma continuous_on_star {s : set E} : continuous_on star s := continuous_star.continuous_on - -lemma continuous_at_star {x : E} : continuous_at star x := continuous_star.continuous_at - -lemma continuous_within_at_star {s : set E} {x : E} : continuous_within_at star s x := -continuous_star.continuous_within_at - -lemma tendsto_star (x : E) : filter.tendsto star (𝓝 x) (𝓝 x⋆) := continuous_star.tendsto x - -lemma filter.tendsto.star {f : α → E} {l : filter α} {y : E} (h : filter.tendsto f l (𝓝 y)) : - filter.tendsto (λ x, (f x)⋆) l (𝓝 y⋆) := -(continuous_star.tendsto y).comp h - -variables [topological_space α] - -lemma continuous.star {f : α → E} (hf : continuous f) : continuous (λ y, star (f y)) := -continuous_star.comp hf - -lemma continuous_at.star {f : α → E} {x : α} (hf : continuous_at f x) : - continuous_at (λ x, (f x)⋆) x := -continuous_at_star.comp hf - -lemma continuous_on.star {f : α → E} {s : set α} (hf : continuous_on f s) : - continuous_on (λ x, (f x)⋆) s := -continuous_star.comp_continuous_on hf - -lemma continuous_within_at.star {f : α → E} {s : set α} {x : α} - (hf : continuous_within_at f s x) : continuous_within_at (λ x, (f x)⋆) s x := hf.star +@[priority 100] +instance normed_star_group.to_has_continuous_star : has_continuous_star E := +⟨star_isometry.continuous⟩ end normed_star_group diff --git a/src/data/complex/is_R_or_C.lean b/src/data/complex/is_R_or_C.lean index 657f93b727416..e417e3e209d13 100644 --- a/src/data/complex/is_R_or_C.lean +++ b/src/data/complex/is_R_or_C.lean @@ -846,7 +846,10 @@ noncomputable def conj_cle : K ≃L[ℝ] K := @conj_lie K _ @[simp, is_R_or_C_simps] lemma conj_cle_norm : ∥(@conj_cle K _ : K →L[ℝ] K)∥ = 1 := (@conj_lie K _).to_linear_isometry.norm_to_continuous_linear_map -@[continuity] lemma continuous_conj : continuous (conj : K → K) := conj_lie.continuous +@[priority 100] +instance : has_continuous_star K := ⟨conj_lie.continuous⟩ + +@[continuity] lemma continuous_conj : continuous (conj : K → K) := continuous_star /-- The `ℝ → K` coercion, as a linear map -/ noncomputable def of_real_am : ℝ →ₐ[ℝ] K := algebra.of_id ℝ K diff --git a/src/topology/algebra/star.lean b/src/topology/algebra/star.lean new file mode 100644 index 0000000000000..f8717f9ed951c --- /dev/null +++ b/src/topology/algebra/star.lean @@ -0,0 +1,86 @@ +/- +Copyright (c) 2022 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import algebra.star.pi +import algebra.star.prod +import topology.algebra.group + +/-! +# Continuity of `star` + +This file defines the `has_continuous_star` typeclass, along with instances on `pi`, `prod`, +`mul_opposite`, and `units`. +-/ + + +open_locale filter topological_space +open filter + +universes u + +variables {ι α R S : Type*} + +/-- Basic hypothesis to talk about a topological space with a continuous `star` operator. -/ +class has_continuous_star (R : Type u) [topological_space R] [has_star R] : Prop := +(continuous_star : continuous (star : R → R)) + +export has_continuous_star (continuous_star) + +section continuity + +variables [topological_space R] [has_star R] [has_continuous_star R] + +lemma continuous_on_star {s : set R} : continuous_on star s := +continuous_star.continuous_on + +lemma continuous_within_at_star {s : set R} {x : R} : continuous_within_at star s x := +continuous_star.continuous_within_at + +lemma continuous_at_star {x : R} : continuous_at star x := +continuous_star.continuous_at + +lemma tendsto_star (a : R) : tendsto star (𝓝 a) (𝓝 (star a)) := +continuous_at_star + +lemma filter.tendsto.star {f : α → R} {l : filter α} {y : R} (h : tendsto f l (𝓝 y)) : + tendsto (λ x, star (f x)) l (𝓝 (star y)) := +(continuous_star.tendsto y).comp h + +variables [topological_space α] {f : α → R} {s : set α} {x : α} + +@[continuity] +lemma continuous.star (hf : continuous f) : continuous (λ x, star (f x)) := +continuous_star.comp hf + +lemma continuous_at.star (hf : continuous_at f x) : continuous_at (λ x, star (f x)) x := +continuous_at_star.comp hf + +lemma continuous_on.star (hf : continuous_on f s) : continuous_on (λ x, star (f x)) s := +continuous_star.comp_continuous_on hf + +lemma continuous_within_at.star (hf : continuous_within_at f s x) : + continuous_within_at (λ x, star (f x)) s x := +hf.star + +end continuity + +section instances + +instance [has_star R] [has_star S] [topological_space R] [topological_space S] + [has_continuous_star R] [has_continuous_star S] : has_continuous_star (R × S) := +⟨(continuous_star.comp continuous_fst).prod_mk (continuous_star.comp continuous_snd)⟩ + +instance {C : ι → Type*} [∀ i, topological_space (C i)] + [∀ i, has_star (C i)] [∀ i, has_continuous_star (C i)] : has_continuous_star (Π i, C i) := +{ continuous_star := continuous_pi (λ i, continuous.star (continuous_apply i)) } + +instance [has_star R] [topological_space R] [has_continuous_star R] : has_continuous_star Rᵐᵒᵖ := +⟨mul_opposite.continuous_op.comp $ mul_opposite.continuous_unop.star⟩ + +instance [monoid R] [star_semigroup R] [topological_space R] [has_continuous_star R] : + has_continuous_star Rˣ := +⟨continuous_induced_rng units.continuous_embed_product.star⟩ + +end instances diff --git a/src/topology/continuous_function/zero_at_infty.lean b/src/topology/continuous_function/zero_at_infty.lean index 0cddc40ff5813..252957f0c6000 100644 --- a/src/topology/continuous_function/zero_at_infty.lean +++ b/src/topology/continuous_function/zero_at_infty.lean @@ -423,15 +423,14 @@ section star /-! ### Star structure It is possible to equip `C₀(α, β)` with a pointwise `star` operation whenever there is a continuous -`star : β → β` for which `star (0 : β) = 0`. However, we have no such minimal type classes (e.g., -`has_continuous_star` or `star_zero_class`) and so the type class assumptions on `β` sufficient to -guarantee these conditions are `[normed_group β]`, `[star_add_monoid β]` and -`[normed_star_group β]`, which allow for the corresponding classes on `C₀(α, β)` essentially -inherited from their counterparts on `α →ᵇ β`. Ultimately, when `β` is a C⋆-ring, then so is -`C₀(α, β)`. +`star : β → β` for which `star (0 : β) = 0`. We don't have quite this weak a typeclass, but +`star_add_monoid` is close enough. + +The `star_add_monoid` and `normed_star_group` classes on `C₀(α, β)` are inherited from their +counterparts on `α →ᵇ β`. Ultimately, when `β` is a C⋆-ring, then so is `C₀(α, β)`. -/ -variables [normed_group β] [star_add_monoid β] [normed_star_group β] +variables [topological_space β] [add_monoid β] [star_add_monoid β] [has_continuous_star β] instance : has_star C₀(α, β) := { star := λ f, @@ -446,20 +445,26 @@ lemma coe_star (f : C₀(α, β)) : ⇑(star f) = star f := rfl lemma star_apply (f : C₀(α, β)) (x : α) : (star f) x = star (f x) := rfl -instance : star_add_monoid C₀(α, β) := +instance [has_continuous_add β] : star_add_monoid C₀(α, β) := { star_involutive := λ f, ext $ λ x, star_star (f x), star_add := λ f g, ext $ λ x, star_add (f x) (g x) } +end star + +section normed_star + +variables [normed_group β] [star_add_monoid β] [normed_star_group β] + instance : normed_star_group C₀(α, β) := { norm_star := λ f, (norm_star f.to_bcf : _) } -end star +end normed_star section star_module -variables {𝕜 : Type*} [semiring 𝕜] [has_star 𝕜] - [normed_group β] [star_add_monoid β] [normed_star_group β] - [module 𝕜 β] [has_continuous_const_smul 𝕜 β] [star_module 𝕜 β] +variables {𝕜 : Type*} [has_zero 𝕜] [has_star 𝕜] + [add_monoid β] [star_add_monoid β] [topological_space β] [has_continuous_star β] + [smul_with_zero 𝕜 β] [has_continuous_const_smul 𝕜 β] [star_module 𝕜 β] instance : star_module 𝕜 C₀(α, β) := { star_smul := λ k f, ext $ λ x, star_smul k (f x) } @@ -468,16 +473,21 @@ end star_module section star_ring -variables [non_unital_normed_ring β] [star_ring β] +variables [non_unital_semiring β] [star_ring β] [topological_space β] [has_continuous_star β] + [topological_semiring β] -instance [normed_star_group β] : star_ring C₀(α, β) := +instance : star_ring C₀(α, β) := { star_mul := λ f g, ext $ λ x, star_mul (f x) (g x), ..zero_at_infty_continuous_map.star_add_monoid } -instance [cstar_ring β] : cstar_ring C₀(α, β) := +end star_ring + +section cstar_ring + +instance [non_unital_normed_ring β] [star_ring β] [cstar_ring β] : cstar_ring C₀(α, β) := { norm_star_mul_self := λ f, @cstar_ring.norm_star_mul_self _ _ _ _ f.to_bcf } -end star_ring +end cstar_ring /-! ### C₀ as a functor diff --git a/src/topology/instances/matrix.lean b/src/topology/instances/matrix.lean index 7a039af9faac6..319ae93cc7f2a 100644 --- a/src/topology/instances/matrix.lean +++ b/src/topology/instances/matrix.lean @@ -4,8 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Oliver Nash, Eric Wieser -/ import linear_algebra.determinant -import topology.algebra.ring import topology.algebra.infinite_sum +import topology.algebra.ring +import topology.algebra.star /-! # Topological properties of matrices @@ -69,13 +70,12 @@ lemma continuous.matrix_transpose {A : X → matrix m n R} (hA : continuous A) : continuous (λ x, (A x)ᵀ) := continuous_matrix $ λ i j, hA.matrix_elem j i -/-! TODO: add a `has_continuous_star` typeclass so we can write -``` -lemma continuous.matrix.conj_transpose [has_star R] {A : X → matrix m n R} (hA : continuous A) : - continuous (λ x, (A x)ᴴ) := +lemma continuous.matrix_conj_transpose [has_star R] [has_continuous_star R] {A : X → matrix m n R} + (hA : continuous A) : continuous (λ x, (A x)ᴴ) := hA.matrix_transpose.matrix_map continuous_star -``` --/ + +instance [has_star R] [has_continuous_star R] : has_continuous_star (matrix m m R) := +⟨continuous_id.matrix_conj_transpose⟩ @[continuity] lemma continuous.matrix_col {A : X → n → R} (hA : continuous A) : continuous (λ x, col (A x)) := @@ -265,6 +265,31 @@ begin rw [tsum_eq_zero_of_not_summable hf, tsum_eq_zero_of_not_summable hft, transpose_zero] }, end +lemma has_sum.matrix_conj_transpose [star_add_monoid R] [has_continuous_star R] + {f : X → matrix m n R} {a : matrix m n R} (hf : has_sum f a) : + has_sum (λ x, (f x)ᴴ) aᴴ := +(hf.map (@matrix.conj_transpose_add_equiv m n R _ _) continuous_id.matrix_conj_transpose : _) + +lemma summable.matrix_conj_transpose [star_add_monoid R] [has_continuous_star R] + {f : X → matrix m n R} (hf : summable f) : + summable (λ x, (f x)ᴴ) := +hf.has_sum.matrix_conj_transpose.summable + +@[simp] lemma summable_matrix_conj_transpose [star_add_monoid R] [has_continuous_star R] + {f : X → matrix m n R} : + summable (λ x, (f x)ᴴ) ↔ summable f := +(summable.map_iff_of_equiv (@matrix.conj_transpose_add_equiv m n R _ _) + (@continuous_id (matrix m n R) _).matrix_conj_transpose (continuous_id.matrix_conj_transpose) : _) + +lemma matrix.conj_transpose_tsum [star_add_monoid R] [has_continuous_star R] [t2_space R] + {f : X → matrix m n R} : (∑' x, f x)ᴴ = ∑' x, (f x)ᴴ := +begin + by_cases hf : summable f, + { exact hf.has_sum.matrix_conj_transpose.tsum_eq.symm }, + { have hft := summable_matrix_conj_transpose.not.mpr hf, + rw [tsum_eq_zero_of_not_summable hf, tsum_eq_zero_of_not_summable hft, conj_transpose_zero] }, +end + lemma has_sum.matrix_diagonal [decidable_eq n] {f : X → n → R} {a : n → R} (hf : has_sum f a) : has_sum (λ x, diagonal (f x)) (diagonal a) := (hf.map (diagonal_add_monoid_hom n R) $ continuous.matrix_diagonal $ by exact continuous_id : _) diff --git a/src/topology/instances/real.lean b/src/topology/instances/real.lean index 2756c58503f1d..515ecfb0f9c1b 100644 --- a/src/topology/instances/real.lean +++ b/src/topology/instances/real.lean @@ -6,6 +6,7 @@ Authors: Johannes Hölzl, Mario Carneiro import topology.metric_space.basic import topology.algebra.uniform_group import topology.algebra.ring +import topology.algebra.star import ring_theory.subring.basic import group_theory.archimedean import algebra.periodic @@ -35,6 +36,8 @@ theorem real.uniform_continuous_neg : uniform_continuous (@has_neg.neg ℝ _) := metric.uniform_continuous_iff.2 $ λ ε ε0, ⟨_, ε0, λ a b h, by rw dist_comm at h; simpa [real.dist_eq] using h⟩ +instance : has_continuous_star ℝ := ⟨continuous_id⟩ + instance : uniform_add_group ℝ := uniform_add_group.mk' real.uniform_continuous_add real.uniform_continuous_neg From 1afdaf96bfd4b62658e486541ee2662615a3b157 Mon Sep 17 00:00:00 2001 From: antoinelab01 Date: Wed, 4 May 2022 07:50:51 +0000 Subject: [PATCH 082/153] feat(linear_algebra/trace): more general versions of `trace_mul_comm` and `trace_conj` (#13874) --- src/linear_algebra/bilinear_map.lean | 6 +++ src/linear_algebra/contraction.lean | 16 ++++++-- src/linear_algebra/trace.lean | 56 +++++++++++++++++++++------- 3 files changed, 62 insertions(+), 16 deletions(-) diff --git a/src/linear_algebra/bilinear_map.lean b/src/linear_algebra/bilinear_map.lean index de7cccddacdca..84cd7d4333f22 100644 --- a/src/linear_algebra/bilinear_map.lean +++ b/src/linear_algebra/bilinear_map.lean @@ -216,6 +216,9 @@ variables {R Pₗ} @[simp] theorem lcomp_apply (f : M →ₗ[R] Nₗ) (g : Nₗ →ₗ[R] Pₗ) (x : M) : lcomp R Pₗ f g x = g (f x) := rfl +theorem lcomp_apply' (f : M →ₗ[R] Nₗ) (g : Nₗ →ₗ[R] Pₗ) : + lcomp R Pₗ f g = g ∘ₗ f := rfl + variables (P σ₂₃) /-- Composing a semilinear map `M → N` and a semilinear map `N → P` to form a semilinear map `M → P` is itself a linear map. -/ @@ -239,6 +242,9 @@ variables {R M Nₗ Pₗ} section @[simp] theorem llcomp_apply (f : Nₗ →ₗ[R] Pₗ) (g : M →ₗ[R] Nₗ) (x : M) : llcomp R M Nₗ Pₗ f g x = f (g x) := rfl + +theorem llcomp_apply' (f : Nₗ →ₗ[R] Pₗ) (g : M →ₗ[R] Nₗ) : + llcomp R M Nₗ Pₗ f g = f ∘ₗ g := rfl end /-- Composing a linear map `Q → N` and a bilinear map `M → N → P` to diff --git a/src/linear_algebra/contraction.lean b/src/linear_algebra/contraction.lean index bb9ded19d814f..423c0d771bab4 100644 --- a/src/linear_algebra/contraction.lean +++ b/src/linear_algebra/contraction.lean @@ -25,11 +25,11 @@ section contraction open tensor_product linear_map matrix open_locale tensor_product big_operators -variables (R M N : Type*) [add_comm_group M] [add_comm_group N] +variables (R M N P : Type*) [add_comm_group M] [add_comm_group N] [add_comm_group P] section comm_ring -variables [comm_ring R] [module R M] [module R N] +variables [comm_ring R] [module R M] [module R N] [module R P] variables {ι : Type*} [decidable_eq ι] [fintype ι] (b : basis ι R M) /-- The natural left-handed pairing between a module and its dual. -/ @@ -44,7 +44,7 @@ def dual_tensor_hom : (module.dual R M) ⊗ N →ₗ[R] M →ₗ[R] N := let M' := module.dual R M in (uncurry R M' N (M →ₗ[R] N) : _ → M' ⊗ N →ₗ[R] M →ₗ[R] N) linear_map.smul_rightₗ -variables {R M N} +variables {R M N P} @[simp] lemma contract_left_apply (f : module.dual R M) (m : M) : contract_left R M (f ⊗ₜ m) = f m := by apply uncurry_apply @@ -56,6 +56,14 @@ variables {R M N} dual_tensor_hom R M N (f ⊗ₜ n) m = (f m) • n := by { dunfold dual_tensor_hom, rw uncurry_apply, refl, } +@[simp] lemma comp_dual_tensor_hom (f : module.dual R M) (n : N) (g : module.dual R N) (p : P) : + (dual_tensor_hom R N P (g ⊗ₜ[R] p)) ∘ₗ (dual_tensor_hom R M N (f ⊗ₜ[R] n)) = + g n • dual_tensor_hom R M P (f ⊗ₜ p) := +begin + ext m, simp only [coe_comp, function.comp_app, dual_tensor_hom_apply, map_smulₛₗ, + ring_hom.id_apply, smul_apply], rw smul_comm, +end + /-- As a matrix, `dual_tensor_hom` evaluated on a basis element of `M* ⊗ N` is a matrix with a single one and zeros elsewhere -/ theorem to_matrix_dual_tensor_hom @@ -98,6 +106,8 @@ end) dual_tensor_hom R M N := rfl +variables (R M N) + variables [module.free R M] [module.finite R M] [nontrivial R] open_locale classical diff --git a/src/linear_algebra/trace.lean b/src/linear_algebra/trace.lean index d6f9c100b8042..ddc990b390689 100644 --- a/src/linear_algebra/trace.lean +++ b/src/linear_algebra/trace.lean @@ -32,6 +32,8 @@ open_locale big_operators open_locale matrix open finite_dimensional +open_locale tensor_product + section variables (R : Type u) [comm_ring R] {M : Type v} [add_comm_group M] [module R M] variables {ι : Type w} [decidable_eq ι] [fintype ι] @@ -103,18 +105,13 @@ theorem trace_conj (g : M →ₗ[R] M) (f : (M →ₗ[R] M)ˣ) : trace R M (↑f * g * ↑f⁻¹) = trace R M g := by { rw trace_mul_comm, simp } -/-- The trace of an endomorphism is invariant under conjugation -/ -@[simp] -theorem trace_conj' (f g : M →ₗ[R] M) [invertible f] : - trace R M (f * g * ⅟ f) = trace R M g := -by { rw trace_mul_comm, simp } - end section -variables (R : Type u) [comm_ring R] {M : Type v} [add_comm_group M] [module R M] -variables {ι : Type w} [fintype ι] +variables {R : Type*} [comm_ring R] {M : Type*} [add_comm_group M] [module R M] +variables (N : Type*) [add_comm_group N] [module R N] +variables {ι : Type*} [fintype ι] /-- The trace of a linear map correspond to the contraction pairing under the isomorphism `End(M) ≃ M* ⊗ M`-/ @@ -137,15 +134,20 @@ end lemma trace_eq_contract_of_basis' [decidable_eq ι] (b : basis ι R M) : (linear_map.trace R M) = (contract_left R M) ∘ₗ (dual_tensor_hom_equiv_of_basis b).symm.to_linear_map := -by simp [linear_equiv.eq_comp_to_linear_map_symm, trace_eq_contract_of_basis R b] +by simp [linear_equiv.eq_comp_to_linear_map_symm, trace_eq_contract_of_basis b] -variables [module.free R M] [module.finite R M] [nontrivial R] +variables (R M) +variables [module.free R M] [module.finite R M] [module.free R N] [module.finite R N] [nontrivial R] /-- When `M` is finite free, the trace of a linear map correspond to the contraction pairing under the isomorphism `End(M) ≃ M* ⊗ M`-/ @[simp] theorem trace_eq_contract : (linear_map.trace R M) ∘ₗ (dual_tensor_hom R M M) = contract_left R M := -trace_eq_contract_of_basis R (module.free.choose_basis R M) +trace_eq_contract_of_basis (module.free.choose_basis R M) + +@[simp] theorem trace_eq_contract_apply (x : module.dual R M ⊗[R] M) : + (linear_map.trace R M) ((dual_tensor_hom R M M) x) = contract_left R M x := +by rw [←comp_apply, trace_eq_contract] open_locale classical @@ -153,8 +155,8 @@ open_locale classical the isomorphism `End(M) ≃ M* ⊗ M`-/ theorem trace_eq_contract' : (linear_map.trace R M) = - (contract_left R M) ∘ₗ (dual_tensor_hom_equiv).symm.to_linear_map := -trace_eq_contract_of_basis' R (module.free.choose_basis R M) + (contract_left R M) ∘ₗ (dual_tensor_hom_equiv R M M).symm.to_linear_map := +trace_eq_contract_of_basis' (module.free.choose_basis R M) /-- The trace of the identity endomorphism is the dimension of the free module -/ @[simp] theorem trace_one : trace R M 1 = (finrank R M : R) := @@ -164,6 +166,34 @@ begin simp, end +variables (M) + +theorem trace_comp_comm : + compr₂ (llcomp R M N M) (trace R M) = compr₂ (llcomp R N M N).flip (trace R N) := +begin + apply (compl₁₂_inj + (dual_tensor_hom_equiv R N M).surjective (dual_tensor_hom_equiv R M N).surjective).1, + ext g m f n, + simp only [tensor_product.algebra_tensor_module.curry_apply, to_fun_eq_coe, + tensor_product.curry_apply, coe_restrict_scalars_eq_coe, compl₁₂_apply, compr₂_apply, flip_apply, + llcomp_apply', comp_dual_tensor_hom, map_smul, trace_eq_contract_apply, contract_left_apply, + smul_eq_mul, mul_comm], +end + +variables {R M} + +theorem trace_comp_comm' (f : M →ₗ[R] N) (g : N →ₗ[R] M) : + trace R M (g ∘ₗ f) = trace R N (f ∘ₗ g) := +begin + have h := ext_iff.1 (ext_iff.1 (trace_comp_comm R M N) g) f, + simp only [llcomp_apply', compr₂_apply, flip_apply] at h, + exact h, +end + +@[simp] theorem trace_conj' (f : M →ₗ[R] M) (e : M ≃ₗ[R] N) : trace R N (e.conj f) = trace R M f := +by rw [e.conj_apply, trace_comp_comm', ←comp_assoc, linear_equiv.comp_coe, + linear_equiv.self_trans_symm, linear_equiv.refl_to_linear_map, id_comp] + end end linear_map From d537897a924791827325d44ece505b10d7a3096e Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 4 May 2022 07:50:52 +0000 Subject: [PATCH 083/153] feat(category_theory/simple): simple objects are indecomposable (#13882) Remarkably tedious. Co-authored-by: Scott Morrison --- .../limits/shapes/biproducts.lean | 29 +++++++++- .../limits/shapes/zero_morphisms.lean | 57 +++++++++++++++++++ src/category_theory/simple.lean | 56 ++++++++++++++++++ 3 files changed, 141 insertions(+), 1 deletion(-) diff --git a/src/category_theory/limits/shapes/biproducts.lean b/src/category_theory/limits/shapes/biproducts.lean index 7a7350aa82694..91e712e77d378 100644 --- a/src/category_theory/limits/shapes/biproducts.lean +++ b/src/category_theory/limits/shapes/biproducts.lean @@ -1068,6 +1068,26 @@ def biprod.unique_up_to_iso (X Y : C) [has_binary_biproduct X Y] {b : binary_bic inv_hom_id' := by rw [← biprod.cone_point_unique_up_to_iso_hom X Y hb, ← biprod.cone_point_unique_up_to_iso_inv X Y hb, iso.inv_hom_id] } +section +variables (X Y : C) [has_binary_biproduct X Y] + +-- There are three further variations, +-- about `is_iso biprod.inr`, `is_iso biprod.fst` and `is_iso biprod.snd`, +-- but any one suffices to prove `indecomposable_of_simple` +-- and they are likely not separately useful. +lemma biprod.is_iso_inl_iff_id_eq_fst_comp_inl : + is_iso (biprod.inl : X ⟶ X ⊞ Y) ↔ 𝟙 (X ⊞ Y) = biprod.fst ≫ biprod.inl := +begin + split, + { introI h, + have := (cancel_epi (inv biprod.inl : X ⊞ Y ⟶ X)).2 biprod.inl_fst, + rw [is_iso.inv_hom_id_assoc, category.comp_id] at this, + rw [this, is_iso.inv_hom_id], }, + { intro h, exact ⟨⟨biprod.fst, biprod.inl_fst, h.symm⟩⟩, }, +end + +end + section biprod_kernel variables (X Y : C) [has_binary_biproduct X Y] @@ -1570,7 +1590,7 @@ begin split_epi_of_idempotent_of_is_colimit_cofork_section_], dsimp only [binary_bicone_of_split_mono_of_cokernel_X], rw [is_colimit_cofork_of_cokernel_cofork_desc, is_cokernel_epi_comp_desc], - simp only [binary_bicone_of_split_mono_of_cokernel_inl, cofork.is_colimit.π_comp_desc, + simp only [binary_bicone_of_split_mono_of_cokernel_inl, cofork.is_colimit.π_comp_desc, cokernel_cofork_of_cofork_π, cofork.π_of_π, add_sub_cancel'_right] end @@ -1645,6 +1665,7 @@ end limits open category_theory.limits +section local attribute [ext] preadditive /-- The existence of binary biproducts implies that there is at most one preadditive structure. -/ @@ -1661,5 +1682,11 @@ begin congr' 2; exact subsingleton.elim _ _ end +end + +variables {C : Type u} [category.{v} C] [has_zero_morphisms C] [has_binary_biproducts C] + +/-- An object is indecomposable if it cannot be written as the biproduct of two nonzero objects. -/ +def indecomposable (X : C) : Prop := ¬ is_zero X ∧ ∀ Y Z, (X ≅ Y ⊞ Z) → is_zero Y ∨ is_zero Z end category_theory diff --git a/src/category_theory/limits/shapes/zero_morphisms.lean b/src/category_theory/limits/shapes/zero_morphisms.lean index a3d1ac5a5a589..11e19b2693bab 100644 --- a/src/category_theory/limits/shapes/zero_morphisms.lean +++ b/src/category_theory/limits/shapes/zero_morphisms.lean @@ -132,6 +132,63 @@ instance : has_zero_morphisms (C ⥤ D) := end +namespace is_zero +variables [has_zero_morphisms C] + +lemma eq_zero_of_src {X Y : C} (o : is_zero X) (f : X ⟶ Y) : f = 0 := +o.eq_of_src _ _ + +lemma eq_zero_of_tgt {X Y : C} (o : is_zero Y) (f : X ⟶ Y) : f = 0 := +o.eq_of_tgt _ _ + +lemma iff_id_eq_zero (X : C) : is_zero X ↔ (𝟙 X = 0) := +⟨λ h, h.eq_of_src _ _, + λ h, ⟨ + λ Y, ⟨⟨⟨0⟩, λ f, by { rw [←id_comp f, ←id_comp default, h, zero_comp, zero_comp], }⟩⟩, + λ Y, ⟨⟨⟨0⟩, λ f, by { rw [←comp_id f, ←comp_id default, h, comp_zero, comp_zero], }⟩⟩⟩⟩ + +lemma of_mono_zero (X Y : C) [mono (0 : X ⟶ Y)] : is_zero X := +(iff_id_eq_zero X).mpr ((cancel_mono (0 : X ⟶ Y)).1 (by simp)) + +lemma of_epi_zero (X Y : C) [epi (0 : X ⟶ Y)] : is_zero Y := +(iff_id_eq_zero Y).mpr ((cancel_epi (0 : X ⟶ Y)).1 (by simp)) + +lemma of_mono_eq_zero {X Y : C} (f : X ⟶ Y) [mono f] (h : f = 0) : is_zero X := +by { unfreezingI { subst h, }, apply of_mono_zero X Y, } + +lemma of_epi_eq_zero {X Y : C} (f : X ⟶ Y) [epi f] (h : f = 0) : is_zero Y := +by { unfreezingI { subst h, }, apply of_epi_zero X Y, } + +lemma iff_split_mono_eq_zero {X Y : C} (f : X ⟶ Y) [split_mono f] : is_zero X ↔ f = 0 := +begin + rw iff_id_eq_zero, + split, + { intro h, rw [←category.id_comp f, h, zero_comp], }, + { intro h, rw [←split_mono.id f], simp [h], }, +end + +lemma iff_split_epi_eq_zero {X Y : C} (f : X ⟶ Y) [split_epi f] : is_zero Y ↔ f = 0 := +begin + rw iff_id_eq_zero, + split, + { intro h, rw [←category.comp_id f, h, comp_zero], }, + { intro h, rw [←split_epi.id f], simp [h], }, +end + +lemma of_mono {X Y : C} (f : X ⟶ Y) [mono f] (i : is_zero Y) : is_zero X := +begin + unfreezingI { have hf := i.eq_zero_of_tgt f, subst hf, }, + exact is_zero.of_mono_zero X Y, +end + +lemma of_epi {X Y : C} (f : X ⟶ Y) [epi f] (i : is_zero X) : is_zero Y := +begin + unfreezingI { have hf := i.eq_zero_of_src f, subst hf, }, + exact is_zero.of_epi_zero X Y, +end + +end is_zero + /-- A category with a zero object has zero morphisms. It is rarely a good idea to use this. Many categories that have a zero object have zero diff --git a/src/category_theory/simple.lean b/src/category_theory/simple.lean index 3f5a8fe86a69d..a19205474b3bc 100644 --- a/src/category_theory/simple.lean +++ b/src/category_theory/simple.lean @@ -16,6 +16,8 @@ is either an isomorphism or zero (but not both). This is formalized as a `Prop` valued typeclass `simple X`. +In some contexts, especially representation theory, simple objects are called "irreducibles". + If a morphism `f` out of a simple object is nonzero and has a kernel, then that kernel is zero. (We state this as `kernel.ι f = 0`, but should add `kernel f ≅ 0`.) @@ -23,6 +25,8 @@ When the category is abelian, being simple is the same as being cosimple (althou state a separate typeclass for this). As a consequence, any nonzero epimorphism out of a simple object is an isomorphism, and any nonzero morphism into a simple object has trivial cokernel. + +We show that any simple object is indecomposable. -/ noncomputable theory @@ -46,6 +50,25 @@ lemma is_iso_of_mono_of_nonzero {X Y : C} [simple Y] {f : X ⟶ Y} [mono f] (w : is_iso f := (simple.mono_is_iso_iff_nonzero f).mpr w +lemma simple.of_iso {X Y : C} [simple Y] (i : X ≅ Y) : simple X := +{ mono_is_iso_iff_nonzero := λ Z f m, begin + resetI, + haveI : mono (f ≫ i.hom) := mono_comp _ _, + split, + { introsI h w, + haveI j : is_iso (f ≫ i.hom), apply_instance, + rw simple.mono_is_iso_iff_nonzero at j, + unfreezingI { subst w, }, + simpa using j, }, + { intro h, + haveI j : is_iso (f ≫ i.hom), + { apply is_iso_of_mono_of_nonzero, + intro w, apply h, + simpa using (cancel_mono i.inv).2 w, }, + rw [←category.comp_id f, ←i.hom_inv_id, ←category.assoc], + apply_instance, }, + end } + lemma kernel_zero_of_nonzero_from_simple {X Y : C} [simple X] {f : X ⟶ Y} [has_kernel f] (w : f ≠ 0) : kernel.ι f = 0 := @@ -71,6 +94,10 @@ instance (X : C) [simple.{v} X] : nontrivial (End X) := nontrivial_of_ne 1 0 (id_nonzero X) section + +lemma simple.not_is_zero (X : C) [simple X] : ¬ is_zero X := +by simpa [limits.is_zero.iff_id_eq_zero] using id_nonzero X + variable [has_zero_object C] open_locale zero_object @@ -136,4 +163,33 @@ end end abelian +section +variables [preadditive C] [has_binary_biproducts C] + +-- There are another three potential variations of this lemma, +-- but as any one suffices to prove `indecomposable_of_simple` we will not give them all. +lemma biprod.is_iso_inl_iff_is_zero (X Y : C) : is_iso (biprod.inl : X ⟶ X ⊞ Y) ↔ is_zero Y := +begin + rw [biprod.is_iso_inl_iff_id_eq_fst_comp_inl, ←biprod.total, add_right_eq_self], + split, + { intro h, replace h := h =≫ biprod.snd, + simpa [←is_zero.iff_split_epi_eq_zero (biprod.snd : X ⊞ Y ⟶ Y)] using h, }, + { intro h, rw is_zero.iff_split_epi_eq_zero (biprod.snd : X ⊞ Y ⟶ Y) at h, rw [h, zero_comp], }, +end + +/-- Any simple object in a preadditive category is indecomposable. -/ +lemma indecomposable_of_simple (X : C) [simple X] : indecomposable X := +⟨simple.not_is_zero X, +λ Y Z i, begin + refine or_iff_not_imp_left.mpr (λ h, _), + rw is_zero.iff_split_mono_eq_zero (biprod.inl : Y ⟶ Y ⊞ Z) at h, + change biprod.inl ≠ 0 at h, + rw ←(simple.mono_is_iso_iff_nonzero biprod.inl) at h, + { rwa biprod.is_iso_inl_iff_is_zero at h, }, + { exact simple.of_iso i.symm, }, + { apply_instance, }, +end⟩ + +end + end category_theory From 269bc85dfd5b866576c9b29af87b4cbf6bd2ab2c Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 4 May 2022 07:50:53 +0000 Subject: [PATCH 084/153] feat(analysis/matrix): add `frobenius_norm_conj_transpose` (#13883) This also moves the existing lemmas about the elementwise norm to the same file. --- src/analysis/matrix.lean | 39 +++++++++++++++++----- src/analysis/normed_space/star/matrix.lean | 16 --------- 2 files changed, 30 insertions(+), 25 deletions(-) diff --git a/src/analysis/matrix.lean b/src/analysis/matrix.lean index 750edd6c17c12..82b0140bd377f 100644 --- a/src/analysis/matrix.lean +++ b/src/analysis/matrix.lean @@ -88,10 +88,6 @@ lemma nnnorm_entry_le_entrywise_sup_nnnorm (A : matrix m n α) {i : m} {j : n} : ∥A i j∥₊ ≤ ∥A∥₊ := (nnnorm_le_pi_nnnorm (A i) j).trans (nnnorm_le_pi_nnnorm A i) -@[simp] lemma nnnorm_transpose (A : matrix m n α) : ∥Aᵀ∥₊ = ∥A∥₊ := -by { simp_rw [pi.nnnorm_def], exact finset.sup_comm _ _ _ } -@[simp] lemma norm_transpose (A : matrix m n α) : ∥Aᵀ∥ = ∥A∥ := congr_arg coe $ nnnorm_transpose A - @[simp] lemma nnnorm_map_eq (A : matrix m n α) (f : α → β) (hf : ∀ a, ∥f a∥₊ = ∥a∥₊) : ∥A.map f∥₊ = ∥A∥₊ := by simp_rw [pi.nnnorm_def, matrix.map, hf] @@ -99,6 +95,20 @@ by simp_rw [pi.nnnorm_def, matrix.map, hf] ∥A.map f∥ = ∥A∥ := (congr_arg (coe : ℝ≥0 → ℝ) $ nnnorm_map_eq A f $ λ a, subtype.ext $ hf a : _) +@[simp] lemma nnnorm_transpose (A : matrix m n α) : ∥Aᵀ∥₊ = ∥A∥₊ := +by { simp_rw [pi.nnnorm_def], exact finset.sup_comm _ _ _ } +@[simp] lemma norm_transpose (A : matrix m n α) : ∥Aᵀ∥ = ∥A∥ := congr_arg coe $ nnnorm_transpose A + +@[simp] lemma nnnorm_conj_transpose [star_add_monoid α] [normed_star_group α] (A : matrix m n α) : + ∥Aᴴ∥₊ = ∥A∥₊ := +(nnnorm_map_eq _ _ nnnorm_star).trans A.nnnorm_transpose +@[simp] lemma norm_conj_transpose [star_add_monoid α] [normed_star_group α] (A : matrix m n α) : + ∥Aᴴ∥ = ∥A∥ := +congr_arg coe $ nnnorm_conj_transpose A + +instance [star_add_monoid α] [normed_star_group α] : normed_star_group (matrix m m α) := +⟨norm_conj_transpose⟩ + @[simp] lemma nnnorm_col (v : m → α) : ∥col v∥₊ = ∥v∥₊ := by simp [pi.nnnorm_def] @[simp] lemma norm_col (v : m → α) : ∥col v∥ = ∥v∥ := congr_arg coe $ nnnorm_col v @@ -360,11 +370,6 @@ lemma frobenius_norm_def (A : matrix m n α) : ∥A∥ = (∑ i j, ∥A i j∥ ^ (2 : ℝ)) ^ (1/2 : ℝ) := (congr_arg coe (frobenius_nnnorm_def A)).trans $ by simp [nnreal.coe_sum] -@[simp] lemma frobenius_nnnorm_transpose (A : matrix m n α) : ∥Aᵀ∥₊ = ∥A∥₊ := -by { rw [frobenius_nnnorm_def, frobenius_nnnorm_def, finset.sum_comm], refl } -@[simp] lemma frobenius_norm_transpose (A : matrix m n α) : ∥Aᵀ∥ = ∥A∥ := -congr_arg coe $ frobenius_nnnorm_transpose A - @[simp] lemma frobenius_nnnorm_map_eq (A : matrix m n α) (f : α → β) (hf : ∀ a, ∥f a∥₊ = ∥a∥₊) : ∥A.map f∥₊ = ∥A∥₊ := by simp_rw [frobenius_nnnorm_def, matrix.map, hf] @@ -372,6 +377,22 @@ by simp_rw [frobenius_nnnorm_def, matrix.map, hf] ∥A.map f∥ = ∥A∥ := (congr_arg (coe : ℝ≥0 → ℝ) $ frobenius_nnnorm_map_eq A f $ λ a, subtype.ext $ hf a : _) +@[simp] lemma frobenius_nnnorm_transpose (A : matrix m n α) : ∥Aᵀ∥₊ = ∥A∥₊ := +by { rw [frobenius_nnnorm_def, frobenius_nnnorm_def, finset.sum_comm], refl } +@[simp] lemma frobenius_norm_transpose (A : matrix m n α) : ∥Aᵀ∥ = ∥A∥ := +congr_arg coe $ frobenius_nnnorm_transpose A + +@[simp] lemma frobenius_nnnorm_conj_transpose [star_add_monoid α] [normed_star_group α] + (A : matrix m n α) : ∥Aᴴ∥₊ = ∥A∥₊ := +(frobenius_nnnorm_map_eq _ _ nnnorm_star).trans A.frobenius_nnnorm_transpose +@[simp] lemma frobenius_norm_conj_transpose [star_add_monoid α] [normed_star_group α] + (A : matrix m n α) : ∥Aᴴ∥ = ∥A∥ := +congr_arg coe $ frobenius_nnnorm_conj_transpose A + +instance frobenius_normed_star_group [star_add_monoid α] [normed_star_group α] : + normed_star_group (matrix m m α) := +⟨frobenius_norm_conj_transpose⟩ + @[simp] lemma frobenius_norm_row (v : m → α) : ∥row v∥ = ∥(pi_Lp.equiv 2 _).symm v∥ := by { rw [frobenius_norm_def, fintype.sum_unique], refl } @[simp] lemma frobenius_nnnorm_row (v : m → α) : ∥row v∥₊ = ∥(pi_Lp.equiv 2 _).symm v∥₊ := diff --git a/src/analysis/normed_space/star/matrix.lean b/src/analysis/normed_space/star/matrix.lean index ec8f115cf4902..356d1b43b9431 100644 --- a/src/analysis/normed_space/star/matrix.lean +++ b/src/analysis/normed_space/star/matrix.lean @@ -18,22 +18,6 @@ open_locale big_operators matrix variables {𝕜 m n E : Type*} -namespace matrix -variables [fintype m] [fintype n] [semi_normed_group E] [star_add_monoid E] [normed_star_group E] - -local attribute [instance] matrix.semi_normed_group - -@[simp] lemma norm_conj_transpose (M : matrix m n E) : ∥Mᴴ∥ = ∥M∥ := -(norm_map_eq _ _ norm_star).trans M.norm_transpose - -@[simp] lemma nnnorm_conj_transpose (M : matrix m n E) : ∥Mᴴ∥₊ = ∥M∥₊ := -subtype.ext M.norm_conj_transpose - -instance : normed_star_group (matrix n n E) := -⟨matrix.norm_conj_transpose⟩ - -end matrix - section entrywise_sup_norm variables [is_R_or_C 𝕜] [fintype n] [decidable_eq n] From 6d370064d8fc595aab91befef52f950c3ab8630e Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 4 May 2022 07:50:54 +0000 Subject: [PATCH 085/153] feat(data/list/basic): add `list.cons_diff` (#13892) --- src/data/list/basic.lean | 18 ++++++++++++++++++ 1 file changed, 18 insertions(+) diff --git a/src/data/list/basic.lean b/src/data/list/basic.lean index 78aa37d0c289f..3e1788223b2b8 100644 --- a/src/data/list/basic.lean +++ b/src/data/list/basic.lean @@ -3365,6 +3365,24 @@ by rw [← diff_cons_right, diff_cons] @[simp] theorem nil_diff (l : list α) : [].diff l = [] := by induction l; [refl, simp only [*, diff_cons, erase_of_not_mem (not_mem_nil _)]] +lemma cons_diff (a : α) (l₁ l₂ : list α) : + (a :: l₁).diff l₂ = if a ∈ l₂ then l₁.diff (l₂.erase a) else a :: l₁.diff l₂ := +begin + induction l₂ with b l₂ ih, { refl }, + rcases eq_or_ne a b with rfl|hne, + { simp }, + { simp only [mem_cons_iff, *, false_or, diff_cons_right], + split_ifs with h₂; simp [diff_erase, list.erase, hne, hne.symm] } +end + +lemma cons_diff_of_mem {a : α} {l₂ : list α} (h : a ∈ l₂) (l₁ : list α) : + (a :: l₁).diff l₂ = l₁.diff (l₂.erase a) := +by rw [cons_diff, if_pos h] + +lemma cons_diff_of_not_mem {a : α} {l₂ : list α} (h : a ∉ l₂) (l₁ : list α) : + (a :: l₁).diff l₂ = a :: l₁.diff l₂ := +by rw [cons_diff, if_neg h] + theorem diff_eq_foldl : ∀ (l₁ l₂ : list α), l₁.diff l₂ = foldl list.erase l₁ l₂ | l₁ [] := rfl | l₁ (a::l₂) := (diff_cons l₁ l₂ a).trans (diff_eq_foldl _ _) From e6b84992199574b59393cf0fbed437d5f0b14ceb Mon Sep 17 00:00:00 2001 From: Adam Topaz Date: Wed, 4 May 2022 07:50:55 +0000 Subject: [PATCH 086/153] feat(ring_theory/valuation/valuation_subring): Adds some equivalent conditions for equivalence of valuations (#13895) Co-authored-by: Scott Morrison --- src/ring_theory/valuation/basic.lean | 52 ++++++++++++++++ .../valuation/valuation_subring.lean | 59 +++++++++++++++++++ 2 files changed, 111 insertions(+) diff --git a/src/ring_theory/valuation/basic.lean b/src/ring_theory/valuation/basic.lean index bb11122d6cee0..4375641751d0c 100644 --- a/src/ring_theory/valuation/basic.lean +++ b/src/ring_theory/valuation/basic.lean @@ -254,6 +254,18 @@ begin rwa [add_comm, max_comm] at h' } end +lemma map_add_eq_of_lt_right (h : v x < v y) : v (x + y) = v y := +begin + convert v.map_add_of_distinct_val _, + { symmetry, rw max_eq_right_iff, exact le_of_lt h }, + { exact ne_of_lt h } +end + +lemma map_add_eq_of_lt_left (h : v y < v x) : v (x + y) = v x := +begin + rw add_comm, exact map_add_eq_of_lt_right _ h, +end + lemma map_eq_of_sub_lt (h : v (y - x) < v x) : v y = v x := begin have := valuation.map_add_of_distinct_val v (ne_of_gt h).symm, @@ -347,6 +359,46 @@ begin rwa h, }, end +lemma is_equiv_iff_val_le_one + [linear_ordered_comm_group_with_zero Γ₀] + [linear_ordered_comm_group_with_zero Γ'₀] + {K : Type*} [division_ring K] + (v : valuation K Γ₀) (v' : valuation K Γ'₀) : + v.is_equiv v' ↔ ∀ {x : K}, v x ≤ 1 ↔ v' x ≤ 1 := +⟨λ h x, by simpa using h x 1, is_equiv_of_val_le_one _ _⟩ + +lemma is_equiv_iff_val_eq_one + [linear_ordered_comm_group_with_zero Γ₀] + [linear_ordered_comm_group_with_zero Γ'₀] + {K : Type*} [division_ring K] + (v : valuation K Γ₀) (v' : valuation K Γ'₀) : + v.is_equiv v' ↔ ∀ {x : K}, v x = 1 ↔ v' x = 1 := +begin + split, + { intros h x, + simpa using @is_equiv.val_eq _ _ _ _ _ _ v v' h x 1 }, + { intros h, apply is_equiv_of_val_le_one, intros x, + split, + { intros hx, + cases lt_or_eq_of_le hx with hx' hx', + { have : v (1 + x) = 1, + { rw ← v.map_one, apply map_add_eq_of_lt_left, simpa }, + rw h at this, + rw (show x = (-1) + (1 + x), by simp), + refine le_trans (v'.map_add _ _) _, + simp [this] }, + { rw h at hx', exact le_of_eq hx' } }, + { intros hx, + cases lt_or_eq_of_le hx with hx' hx', + { have : v' (1 + x) = 1, + { rw ← v'.map_one, apply map_add_eq_of_lt_left, simpa }, + rw ← h at this, + rw (show x = (-1) + (1 + x), by simp), + refine le_trans (v.map_add _ _) _, + simp [this] }, + { rw ← h at hx', exact le_of_eq hx' } } } +end + end section supp diff --git a/src/ring_theory/valuation/valuation_subring.lean b/src/ring_theory/valuation/valuation_subring.lean index 9da94e1ba2550..83bfc7285a718 100644 --- a/src/ring_theory/valuation/valuation_subring.lean +++ b/src/ring_theory/valuation/valuation_subring.lean @@ -304,3 +304,62 @@ instance linear_order_overring : linear_order { S | A ≤ S } := end order end valuation_subring + +namespace valuation + +variables {K} {Γ Γ₁ Γ₂ : Type*} + [linear_ordered_comm_group_with_zero Γ] + [linear_ordered_comm_group_with_zero Γ₁] + [linear_ordered_comm_group_with_zero Γ₂] + (v : valuation K Γ) + (v₁ : valuation K Γ₁) + (v₂ : valuation K Γ₂) + +/-- The valuation subring associated to a valuation. -/ +def valuation_subring : valuation_subring K := +{ mem_or_inv_mem' := begin + intros x, + cases le_or_lt (v x) 1, + { left, exact h }, + { right, change v x⁻¹ ≤ 1, + rw [v.map_inv, ← inv_one, inv_le_inv₀], + { exact le_of_lt h }, + { intro c, simpa [c] using h }, + { exact one_ne_zero } } + end, + .. v.integer } + +@[simp] +lemma mem_valuation_subring_iff (x : K) : x ∈ v.valuation_subring ↔ v x ≤ 1 := iff.refl _ + +lemma is_equiv_iff_valuation_subring : v₁.is_equiv v₂ ↔ + v₁.valuation_subring = v₂.valuation_subring := +begin + split, + { intros h, ext x, specialize h x 1, simpa using h }, + { intros h, apply is_equiv_of_val_le_one, + intros x, + have : x ∈ v₁.valuation_subring ↔ x ∈ v₂.valuation_subring, by rw h, + simpa using this } +end + +lemma is_equiv_valuation_valuation_subring : + v.is_equiv v.valuation_subring.valuation := +begin + rw [is_equiv_iff_val_le_one], + intro x, + rw valuation_subring.valuation_le_one_iff, + refl, +end + +end valuation + +namespace valuation_subring + +variables {K} (A : valuation_subring K) + +@[simp] +lemma valuation_subring_valuation : A.valuation.valuation_subring = A := +by { ext, rw ← A.valuation_le_one_iff, refl } + +end valuation_subring From 455393d346008ac4b08dc7ae49888d4a1d94e6f4 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Wed, 4 May 2022 07:50:57 +0000 Subject: [PATCH 087/153] refactor(group_theory/{submonoid, subsemigroup}/{center, centralizer}): move set.center and set.centralizer into subsemigroup (#13903) This moves `set.center` and `set.centralizer` (the center and centralizers for a magma) into `group_theory/subsemigroup/{center, centralizer}` so that we can define the center and centralizers for semigroups in #13627. --- src/group_theory/submonoid/center.lean | 104 +--------------- src/group_theory/submonoid/centralizer.lean | 98 +-------------- src/group_theory/subsemigroup/center.lean | 114 ++++++++++++++++++ .../subsemigroup/centralizer.lean | 109 +++++++++++++++++ 4 files changed, 231 insertions(+), 194 deletions(-) create mode 100644 src/group_theory/subsemigroup/center.lean create mode 100644 src/group_theory/subsemigroup/centralizer.lean diff --git a/src/group_theory/submonoid/center.lean b/src/group_theory/submonoid/center.lean index 2649bbec6bae5..725ccc143d9df 100644 --- a/src/group_theory/submonoid/center.lean +++ b/src/group_theory/submonoid/center.lean @@ -4,120 +4,24 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ import group_theory.submonoid.operations +import group_theory.subsemigroup.center import data.fintype.basic /-! -# Centers of magmas and monoids +# Centers of monoids ## Main definitions -* `set.center`: the center of a magma * `submonoid.center`: the center of a monoid -* `set.add_center`: the center of an additive magma * `add_submonoid.center`: the center of an additive monoid We provide `subgroup.center`, `add_subgroup.center`, `subsemiring.center`, and `subring.center` in other files. -/ -variables {M : Type*} - -namespace set - -variables (M) - -/-- The center of a magma. -/ -@[to_additive add_center /-" The center of an additive magma. "-/] -def center [has_mul M] : set M := {z | ∀ m, m * z = z * m} - -@[to_additive mem_add_center] -lemma mem_center_iff [has_mul M] {z : M} : z ∈ center M ↔ ∀ g, g * z = z * g := iff.rfl - -instance decidable_mem_center [has_mul M] [decidable_eq M] [fintype M] : - decidable_pred (∈ center M) := -λ _, decidable_of_iff' _ (mem_center_iff M) - -@[simp, to_additive zero_mem_add_center] -lemma one_mem_center [mul_one_class M] : (1 : M) ∈ set.center M := by simp [mem_center_iff] - -@[simp] -lemma zero_mem_center [mul_zero_class M] : (0 : M) ∈ set.center M := by simp [mem_center_iff] - -variables {M} - -@[simp, to_additive add_mem_add_center] -lemma mul_mem_center [semigroup M] {a b : M} - (ha : a ∈ set.center M) (hb : b ∈ set.center M) : a * b ∈ set.center M := -λ g, by rw [mul_assoc, ←hb g, ← mul_assoc, ha g, mul_assoc] - -@[simp, to_additive neg_mem_add_center] -lemma inv_mem_center [group M] {a : M} (ha : a ∈ set.center M) : a⁻¹ ∈ set.center M := -λ g, by rw [← inv_inj, mul_inv_rev, inv_inv, ← ha, mul_inv_rev, inv_inv] - -@[simp] -lemma add_mem_center [distrib M] {a b : M} - (ha : a ∈ set.center M) (hb : b ∈ set.center M) : a + b ∈ set.center M := -λ c, by rw [add_mul, mul_add, ha c, hb c] - -@[simp] -lemma neg_mem_center [ring M] {a : M} (ha : a ∈ set.center M) : -a ∈ set.center M := -λ c, by rw [←neg_mul_comm, ha (-c), neg_mul_comm] - -@[to_additive subset_add_center_add_units] -lemma subset_center_units [monoid M] : - (coe : Mˣ → M) ⁻¹' center M ⊆ set.center Mˣ := -λ a ha b, units.ext $ ha _ - -lemma center_units_subset [group_with_zero M] : - set.center Mˣ ⊆ (coe : Mˣ → M) ⁻¹' center M := -λ a ha b, begin - obtain rfl | hb := eq_or_ne b 0, - { rw [zero_mul, mul_zero], }, - { exact units.ext_iff.mp (ha (units.mk0 _ hb)) } -end - -/-- In a group with zero, the center of the units is the preimage of the center. -/ -lemma center_units_eq [group_with_zero M] : - set.center Mˣ = (coe : Mˣ → M) ⁻¹' center M := -subset.antisymm center_units_subset subset_center_units - -@[simp] -lemma inv_mem_center₀ [group_with_zero M] {a : M} (ha : a ∈ set.center M) : a⁻¹ ∈ set.center M := -begin - obtain rfl | ha0 := eq_or_ne a 0, - { rw inv_zero, exact zero_mem_center M }, - rcases is_unit.mk0 _ ha0 with ⟨a, rfl⟩, - rw ←units.coe_inv', - exact center_units_subset (inv_mem_center (subset_center_units ha)), -end - -@[simp, to_additive sub_mem_add_center] -lemma div_mem_center [group M] {a b : M} (ha : a ∈ set.center M) (hb : b ∈ set.center M) : - a / b ∈ set.center M := -begin - rw [div_eq_mul_inv], - exact mul_mem_center ha (inv_mem_center hb), -end - -@[simp] -lemma div_mem_center₀ [group_with_zero M] {a b : M} (ha : a ∈ set.center M) - (hb : b ∈ set.center M) : a / b ∈ set.center M := -begin - rw div_eq_mul_inv, - exact mul_mem_center ha (inv_mem_center₀ hb), -end - -variables (M) - -@[simp, to_additive add_center_eq_univ] -lemma center_eq_univ [comm_semigroup M] : center M = set.univ := -subset.antisymm (subset_univ _) $ λ x _ y, mul_comm y x - -end set - namespace submonoid section -variables (M) [monoid M] +variables (M : Type*) [monoid M] /-- The center of a monoid `M` is the set of elements that commute with everything in `M` -/ @[to_additive "The center of a monoid `M` is the set of elements that commute with everything in @@ -144,7 +48,7 @@ instance : comm_monoid (center M) := end section -variables (M) [comm_monoid M] +variables (M : Type*) [comm_monoid M] @[simp] lemma center_eq_top : center M = ⊤ := set_like.coe_injective (set.center_eq_univ M) diff --git a/src/group_theory/submonoid/centralizer.lean b/src/group_theory/submonoid/centralizer.lean index 5fb1c68401a0f..c3bac6cd6511c 100644 --- a/src/group_theory/submonoid/centralizer.lean +++ b/src/group_theory/submonoid/centralizer.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Thomas Browning. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Thomas Browning -/ +import group_theory.subsemigroup.centralizer import group_theory.submonoid.center /-! @@ -10,108 +11,17 @@ import group_theory.submonoid.center ## Main definitions -* `set.centralizer`: the center of a magma -* `submonoid.centralizer`: the center of a monoid -* `set.add_centralizer`: the center of an additive magma -* `add_submonoid.centralizer`: the center of an additive monoid +* `submonoid.centralizer`: the centralizer of a subset of a monoid +* `add_submonoid.centralizer`: the centralizer of a subset of an additive monoid We provide `subgroup.centralizer`, `add_subgroup.centralizer` in other files. -/ variables {M : Type*} {S T : set M} -namespace set - -variables (S) - -/-- The centralizer of a subset of a magma. -/ -@[to_additive add_centralizer /-" The centralizer of a subset of an additive magma. "-/] -def centralizer [has_mul M] : set M := {c | ∀ m ∈ S, m * c = c * m} - -variables {S} - -@[to_additive mem_add_centralizer] -lemma mem_centralizer_iff [has_mul M] {c : M} : c ∈ centralizer S ↔ ∀ m ∈ S, m * c = c * m := -iff.rfl - -@[to_additive decidable_mem_add_centralizer] -instance decidable_mem_centralizer [has_mul M] [decidable_eq M] [fintype M] - [decidable_pred (∈ S)] : decidable_pred (∈ centralizer S) := -λ _, decidable_of_iff' _ (mem_centralizer_iff) - -variables (S) - -@[simp, to_additive zero_mem_add_centralizer] -lemma one_mem_centralizer [mul_one_class M] : (1 : M) ∈ centralizer S := -by simp [mem_centralizer_iff] - -@[simp] -lemma zero_mem_centralizer [mul_zero_class M] : (0 : M) ∈ centralizer S := -by simp [mem_centralizer_iff] - -variables {S} {a b : M} - -@[simp, to_additive add_mem_add_centralizer] -lemma mul_mem_centralizer [semigroup M] (ha : a ∈ centralizer S) (hb : b ∈ centralizer S) : - a * b ∈ centralizer S := -λ g hg, by rw [mul_assoc, ←hb g hg, ← mul_assoc, ha g hg, mul_assoc] - -@[simp, to_additive neg_mem_add_centralizer] -lemma inv_mem_centralizer [group M] (ha : a ∈ centralizer S) : a⁻¹ ∈ centralizer S := -λ g hg, by rw [mul_inv_eq_iff_eq_mul, mul_assoc, eq_inv_mul_iff_mul_eq, ha g hg] - -@[simp] -lemma add_mem_centralizer [distrib M] (ha : a ∈ centralizer S) (hb : b ∈ centralizer S) : - a + b ∈ centralizer S := -λ c hc, by rw [add_mul, mul_add, ha c hc, hb c hc] - -@[simp] -lemma neg_mem_centralizer [has_mul M] [has_distrib_neg M] (ha : a ∈ centralizer S) : - -a ∈ centralizer S := -λ c hc, by rw [mul_neg, ha c hc, neg_mul] - -@[simp] -lemma inv_mem_centralizer₀ [group_with_zero M] (ha : a ∈ centralizer S) : a⁻¹ ∈ centralizer S := -(eq_or_ne a 0).elim (λ h, by { rw [h, inv_zero], exact zero_mem_centralizer S }) - (λ ha0 c hc, by rw [mul_inv_eq_iff_eq_mul₀ ha0, mul_assoc, eq_inv_mul_iff_mul_eq₀ ha0, ha c hc]) - -@[simp, to_additive sub_mem_add_centralizer] -lemma div_mem_centralizer [group M] (ha : a ∈ centralizer S) (hb : b ∈ centralizer S) : - a / b ∈ centralizer S := -begin - rw [div_eq_mul_inv], - exact mul_mem_centralizer ha (inv_mem_centralizer hb), -end - -@[simp] -lemma div_mem_centralizer₀ [group_with_zero M] (ha : a ∈ centralizer S) (hb : b ∈ centralizer S) : - a / b ∈ centralizer S := -begin - rw div_eq_mul_inv, - exact mul_mem_centralizer ha (inv_mem_centralizer₀ hb), -end - -@[to_additive add_centralizer_subset] -lemma centralizer_subset [has_mul M] (h : S ⊆ T) : centralizer T ⊆ centralizer S := -λ t ht s hs, ht s (h hs) - -variables (M) - -@[simp, to_additive add_centralizer_univ] -lemma centralizer_univ [has_mul M] : centralizer univ = center M := -subset.antisymm (λ a ha b, ha b (set.mem_univ b)) (λ a ha b hb, ha b) - -variables {M} (S) - -@[simp, to_additive add_centralizer_eq_univ] -lemma centralizer_eq_univ [comm_semigroup M] : centralizer S = univ := -subset.antisymm (subset_univ _) $ λ x hx y hy, mul_comm y x - -end set - namespace submonoid section -variables {M} [monoid M] (S) +variables [monoid M] (S) /-- The centralizer of a subset of a monoid `M`. -/ @[to_additive "The centralizer of a subset of an additive monoid."] diff --git a/src/group_theory/subsemigroup/center.lean b/src/group_theory/subsemigroup/center.lean new file mode 100644 index 0000000000000..c3f70b658919f --- /dev/null +++ b/src/group_theory/subsemigroup/center.lean @@ -0,0 +1,114 @@ +/- +Copyright (c) 2021 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser, Jireh Loreaux +-/ +import group_theory.subsemigroup.operations +import data.fintype.basic + +/-! +# Centers of magmas and semigroups + +## Main definitions + +* `set.center`: the center of a magma +* `set.add_center`: the center of an additive magma + +We provide `submonoid.center`, `add_submonoid.center`, `subgroup.center`, `add_subgroup.center`, +`subsemiring.center`, and `subring.center` in other files. +-/ + +variables {M : Type*} + +namespace set + +variables (M) + +/-- The center of a magma. -/ +@[to_additive add_center /-" The center of an additive magma. "-/] +def center [has_mul M] : set M := {z | ∀ m, m * z = z * m} + +@[to_additive mem_add_center] +lemma mem_center_iff [has_mul M] {z : M} : z ∈ center M ↔ ∀ g, g * z = z * g := iff.rfl + +instance decidable_mem_center [has_mul M] [decidable_eq M] [fintype M] : + decidable_pred (∈ center M) := +λ _, decidable_of_iff' _ (mem_center_iff M) + +@[simp, to_additive zero_mem_add_center] +lemma one_mem_center [mul_one_class M] : (1 : M) ∈ set.center M := by simp [mem_center_iff] + +@[simp] +lemma zero_mem_center [mul_zero_class M] : (0 : M) ∈ set.center M := by simp [mem_center_iff] + +variables {M} + +@[simp, to_additive add_mem_add_center] +lemma mul_mem_center [semigroup M] {a b : M} + (ha : a ∈ set.center M) (hb : b ∈ set.center M) : a * b ∈ set.center M := +λ g, by rw [mul_assoc, ←hb g, ← mul_assoc, ha g, mul_assoc] + +@[simp, to_additive neg_mem_add_center] +lemma inv_mem_center [group M] {a : M} (ha : a ∈ set.center M) : a⁻¹ ∈ set.center M := +λ g, by rw [← inv_inj, mul_inv_rev, inv_inv, ← ha, mul_inv_rev, inv_inv] + +@[simp] +lemma add_mem_center [distrib M] {a b : M} + (ha : a ∈ set.center M) (hb : b ∈ set.center M) : a + b ∈ set.center M := +λ c, by rw [add_mul, mul_add, ha c, hb c] + +@[simp] +lemma neg_mem_center [ring M] {a : M} (ha : a ∈ set.center M) : -a ∈ set.center M := +λ c, by rw [←neg_mul_comm, ha (-c), neg_mul_comm] + +@[to_additive subset_add_center_add_units] +lemma subset_center_units [monoid M] : + (coe : Mˣ → M) ⁻¹' center M ⊆ set.center Mˣ := +λ a ha b, units.ext $ ha _ + +lemma center_units_subset [group_with_zero M] : + set.center Mˣ ⊆ (coe : Mˣ → M) ⁻¹' center M := +λ a ha b, begin + obtain rfl | hb := eq_or_ne b 0, + { rw [zero_mul, mul_zero], }, + { exact units.ext_iff.mp (ha (units.mk0 _ hb)) } +end + +/-- In a group with zero, the center of the units is the preimage of the center. -/ +lemma center_units_eq [group_with_zero M] : + set.center Mˣ = (coe : Mˣ → M) ⁻¹' center M := +subset.antisymm center_units_subset subset_center_units + +@[simp] +lemma inv_mem_center₀ [group_with_zero M] {a : M} (ha : a ∈ set.center M) : a⁻¹ ∈ set.center M := +begin + obtain rfl | ha0 := eq_or_ne a 0, + { rw inv_zero, exact zero_mem_center M }, + rcases is_unit.mk0 _ ha0 with ⟨a, rfl⟩, + rw ←units.coe_inv', + exact center_units_subset (inv_mem_center (subset_center_units ha)), +end + +@[simp, to_additive sub_mem_add_center] +lemma div_mem_center [group M] {a b : M} (ha : a ∈ set.center M) (hb : b ∈ set.center M) : + a / b ∈ set.center M := +begin + rw [div_eq_mul_inv], + exact mul_mem_center ha (inv_mem_center hb), +end + +@[simp] +lemma div_mem_center₀ [group_with_zero M] {a b : M} (ha : a ∈ set.center M) + (hb : b ∈ set.center M) : a / b ∈ set.center M := +begin + rw div_eq_mul_inv, + exact mul_mem_center ha (inv_mem_center₀ hb), +end + +variables (M) + +@[simp, to_additive add_center_eq_univ] +lemma center_eq_univ [comm_semigroup M] : center M = set.univ := +subset.antisymm (subset_univ _) $ λ x _ y, mul_comm y x + +end set diff --git a/src/group_theory/subsemigroup/centralizer.lean b/src/group_theory/subsemigroup/centralizer.lean new file mode 100644 index 0000000000000..031c3120cc819 --- /dev/null +++ b/src/group_theory/subsemigroup/centralizer.lean @@ -0,0 +1,109 @@ +/- +Copyright (c) 2021 Thomas Browning. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Thomas Browning, Jireh Loreaux +-/ +import group_theory.subsemigroup.center + +/-! +# Centralizers of magmas and semigroups + +## Main definitions + +* `set.centralizer`: the centralizer of a subset of a magma +* `set.add_centralizer`: the centralizer of a subset of an additive magma + +We provide `monoid.centralizer`, `add_monoid.centralizer`, `subgroup.centralizer`, and +`add_subgroup.centralizer` in other files. +-/ + +variables {M : Type*} {S T : set M} + +namespace set + +variables (S) + +/-- The centralizer of a subset of a magma. -/ +@[to_additive add_centralizer /-" The centralizer of a subset of an additive magma. "-/] +def centralizer [has_mul M] : set M := {c | ∀ m ∈ S, m * c = c * m} + +variables {S} + +@[to_additive mem_add_centralizer] +lemma mem_centralizer_iff [has_mul M] {c : M} : c ∈ centralizer S ↔ ∀ m ∈ S, m * c = c * m := +iff.rfl + +@[to_additive decidable_mem_add_centralizer] +instance decidable_mem_centralizer [has_mul M] [decidable_eq M] [fintype M] + [decidable_pred (∈ S)] : decidable_pred (∈ centralizer S) := +λ _, decidable_of_iff' _ (mem_centralizer_iff) + +variables (S) + +@[simp, to_additive zero_mem_add_centralizer] +lemma one_mem_centralizer [mul_one_class M] : (1 : M) ∈ centralizer S := +by simp [mem_centralizer_iff] + +@[simp] +lemma zero_mem_centralizer [mul_zero_class M] : (0 : M) ∈ centralizer S := +by simp [mem_centralizer_iff] + +variables {S} {a b : M} + +@[simp, to_additive add_mem_add_centralizer] +lemma mul_mem_centralizer [semigroup M] (ha : a ∈ centralizer S) (hb : b ∈ centralizer S) : + a * b ∈ centralizer S := +λ g hg, by rw [mul_assoc, ←hb g hg, ← mul_assoc, ha g hg, mul_assoc] + +@[simp, to_additive neg_mem_add_centralizer] +lemma inv_mem_centralizer [group M] (ha : a ∈ centralizer S) : a⁻¹ ∈ centralizer S := +λ g hg, by rw [mul_inv_eq_iff_eq_mul, mul_assoc, eq_inv_mul_iff_mul_eq, ha g hg] + +@[simp] +lemma add_mem_centralizer [distrib M] (ha : a ∈ centralizer S) (hb : b ∈ centralizer S) : + a + b ∈ centralizer S := +λ c hc, by rw [add_mul, mul_add, ha c hc, hb c hc] + +@[simp] +lemma neg_mem_centralizer [has_mul M] [has_distrib_neg M] (ha : a ∈ centralizer S) : + -a ∈ centralizer S := +λ c hc, by rw [mul_neg, ha c hc, neg_mul] + +@[simp] +lemma inv_mem_centralizer₀ [group_with_zero M] (ha : a ∈ centralizer S) : a⁻¹ ∈ centralizer S := +(eq_or_ne a 0).elim (λ h, by { rw [h, inv_zero], exact zero_mem_centralizer S }) + (λ ha0 c hc, by rw [mul_inv_eq_iff_eq_mul₀ ha0, mul_assoc, eq_inv_mul_iff_mul_eq₀ ha0, ha c hc]) + +@[simp, to_additive sub_mem_add_centralizer] +lemma div_mem_centralizer [group M] (ha : a ∈ centralizer S) (hb : b ∈ centralizer S) : + a / b ∈ centralizer S := +begin + rw [div_eq_mul_inv], + exact mul_mem_centralizer ha (inv_mem_centralizer hb), +end + +@[simp] +lemma div_mem_centralizer₀ [group_with_zero M] (ha : a ∈ centralizer S) (hb : b ∈ centralizer S) : + a / b ∈ centralizer S := +begin + rw div_eq_mul_inv, + exact mul_mem_centralizer ha (inv_mem_centralizer₀ hb), +end + +@[to_additive add_centralizer_subset] +lemma centralizer_subset [has_mul M] (h : S ⊆ T) : centralizer T ⊆ centralizer S := +λ t ht s hs, ht s (h hs) + +variables (M) + +@[simp, to_additive add_centralizer_univ] +lemma centralizer_univ [has_mul M] : centralizer univ = center M := +subset.antisymm (λ a ha b, ha b (set.mem_univ b)) (λ a ha b hb, ha b) + +variables {M} (S) + +@[simp, to_additive add_centralizer_eq_univ] +lemma centralizer_eq_univ [comm_semigroup M] : centralizer S = univ := +subset.antisymm (subset_univ _) $ λ x hx y hy, mul_comm y x + +end set From 098ab17f68a4533aae57771d2bfc3058a6165de6 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 4 May 2022 07:50:58 +0000 Subject: [PATCH 088/153] feat(category_theory/simple): nonzero morphisms to/from a simple are epi/mono (#13905) Co-authored-by: Scott Morrison --- src/category_theory/preadditive/schur.lean | 7 ++++++- src/category_theory/simple.lean | 13 +++++++++++++ 2 files changed, 19 insertions(+), 1 deletion(-) diff --git a/src/category_theory/preadditive/schur.lean b/src/category_theory/preadditive/schur.lean index 22566b9338844..e37533d9f1990 100644 --- a/src/category_theory/preadditive/schur.lean +++ b/src/category_theory/preadditive/schur.lean @@ -28,6 +28,11 @@ open category_theory.limits variables {C : Type*} [category C] variables [preadditive C] +-- See also `epi_of_nonzero_to_simple`, which does not require `preadditive C`. +lemma mono_of_nonzero_from_simple [has_kernels C] {X Y : C} [simple X] {f : X ⟶ Y} (w : f ≠ 0) : + mono f := +preadditive.mono_of_kernel_zero (kernel_zero_of_nonzero_from_simple w) + /-- The part of **Schur's lemma** that holds in any preadditive category with kernels: that a nonzero morphism between simple objects is an isomorphism. @@ -35,7 +40,7 @@ that a nonzero morphism between simple objects is an isomorphism. lemma is_iso_of_hom_simple [has_kernels C] {X Y : C} [simple X] [simple Y] {f : X ⟶ Y} (w : f ≠ 0) : is_iso f := begin - haveI : mono f := preadditive.mono_of_kernel_zero (kernel_zero_of_nonzero_from_simple w), + haveI := mono_of_nonzero_from_simple w, exact is_iso_of_mono_of_nonzero w end diff --git a/src/category_theory/simple.lean b/src/category_theory/simple.lean index a19205474b3bc..7d5fa5786e259 100644 --- a/src/category_theory/simple.lean +++ b/src/category_theory/simple.lean @@ -79,6 +79,19 @@ begin exact w (eq_zero_of_epi_kernel f), end +/-- +A nonzero morphism `f` to a simple object is an epimorphism +(assuming `f` has an image, and `C` has equalizers). +-/ +-- See also `mono_of_nonzero_from_simple`, which requires `preadditive C`. +lemma epi_of_nonzero_to_simple [has_equalizers C] {X Y : C} [simple Y] + {f : X ⟶ Y} [has_image f] (w : f ≠ 0) : epi f := +begin + rw ←image.fac f, + haveI : is_iso (image.ι f) := is_iso_of_mono_of_nonzero (λ h, w (eq_zero_of_image_eq_zero h)), + apply epi_comp, +end + lemma mono_to_simple_zero_of_not_iso {X Y : C} [simple Y] {f : X ⟶ Y} [mono f] (w : is_iso f → false) : f = 0 := begin From d0efe252cb8bd36e8c1e5557a97e6c5bf31ba11c Mon Sep 17 00:00:00 2001 From: Bhavik Mehta Date: Wed, 4 May 2022 07:50:59 +0000 Subject: [PATCH 089/153] feat(data/finset/prod): diag of union (#13916) Lemmas about diag and off diag in relation to simple finset constructions. --- src/data/finset/prod.lean | 42 ++++++++++++++++++++++++++++++++++++++- 1 file changed, 41 insertions(+), 1 deletion(-) diff --git a/src/data/finset/prod.lean b/src/data/finset/prod.lean index cb460e3bb7969..9b4c075613496 100644 --- a/src/data/finset/prod.lean +++ b/src/data/finset/prod.lean @@ -139,7 +139,7 @@ by { ext ⟨x, y⟩, simp only [and_or_distrib_left, mem_union, mem_product] } end prod section diag -variables (s : finset α) [decidable_eq α] +variables (s t : finset α) [decidable_eq α] /-- Given a finite set `s`, the diagonal, `s.diag` is the set of pairs of the form `(a, a)` for `a ∈ s`. -/ @@ -183,5 +183,45 @@ filter_union_filter_neg_eq _ _ @[simp] lemma disjoint_diag_off_diag : disjoint s.diag s.off_diag := disjoint_filter_filter_neg _ _ +lemma product_sdiff_diag : s.product s \ s.diag = s.off_diag := +by rw [←diag_union_off_diag, union_comm, union_sdiff_self, + sdiff_eq_self_of_disjoint (disjoint_diag_off_diag _).symm] + +lemma product_sdiff_off_diag : s.product s \ s.off_diag = s.diag := +by rw [←diag_union_off_diag, union_sdiff_self, sdiff_eq_self_of_disjoint (disjoint_diag_off_diag _)] + +lemma diag_union : (s ∪ t).diag = s.diag ∪ t.diag := +by { ext ⟨i, j⟩, simp only [mem_diag, mem_union, or_and_distrib_right] } + +variables {s t} + +lemma off_diag_union (h : disjoint s t) : + (s ∪ t).off_diag = s.off_diag ∪ t.off_diag ∪ s.product t ∪ t.product s := +begin + rw [off_diag, union_product, product_union, product_union, union_comm _ (t.product t), + union_assoc, union_left_comm (s.product t), ←union_assoc, filter_union, filter_union, ←off_diag, + ←off_diag, filter_true_of_mem, ←union_assoc], + simp only [mem_union, mem_product, ne.def, prod.forall], + rintro i j (⟨hi, hj⟩ | ⟨hi, hj⟩), + { exact h.forall_ne_finset hi hj }, + { exact h.symm.forall_ne_finset hi hj }, +end + +variables (a : α) + +@[simp] lemma off_diag_singleton : ({a} : finset α).off_diag = ∅ := +by simp [←finset.card_eq_zero] + +lemma diag_singleton : ({a} : finset α).diag = {(a, a)} := +by rw [←product_sdiff_off_diag, off_diag_singleton, sdiff_empty, singleton_product_singleton] + +lemma diag_insert : (insert a s).diag = insert (a, a) s.diag := +by rw [insert_eq, insert_eq, diag_union, diag_singleton] + +lemma off_diag_insert (has : a ∉ s) : + (insert a s).off_diag = s.off_diag ∪ ({a} : finset α).product s ∪ s.product {a} := +by rw [insert_eq, union_comm, off_diag_union (disjoint_singleton_right.2 has), off_diag_singleton, + union_empty, union_right_comm] + end diag end finset From af4c6c8afa4b7136902f6359bf01c5fd888cdce6 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Wed, 4 May 2022 07:51:00 +0000 Subject: [PATCH 090/153] chore(ring_theory/polynomial/basic): golf polynomial_not_is_field (#13919) --- src/ring_theory/polynomial/basic.lean | 23 ++++++++++------------- 1 file changed, 10 insertions(+), 13 deletions(-) diff --git a/src/ring_theory/polynomial/basic.lean b/src/ring_theory/polynomial/basic.lean index e698de40d457e..aa146e5e2066d 100644 --- a/src/ring_theory/polynomial/basic.lean +++ b/src/ring_theory/polynomial/basic.lean @@ -527,19 +527,16 @@ variables [ring R] /-- `polynomial R` is never a field for any ring `R`. -/ lemma polynomial_not_is_field : ¬ is_field R[X] := begin - by_contradiction hR, - by_cases hR' : ∃ (x y : R), x ≠ y, - { haveI : nontrivial R := let ⟨x, y, hxy⟩ := hR' in nontrivial_of_ne x y hxy, - obtain ⟨p, hp⟩ := hR.mul_inv_cancel X_ne_zero, - by_cases hp0 : p = 0, - { replace hp := congr_arg degree hp, - rw [hp0, mul_zero, degree_zero, degree_one] at hp, - contradiction }, - { have : p.degree < (X * p).degree := (X_mul.symm : p * X = _) ▸ degree_lt_degree_mul_X hp0, - rw [congr_arg degree hp, degree_one, nat.with_bot.lt_zero_iff, degree_eq_bot] at this, - exact hp0 this } }, - { push_neg at hR', - exact let ⟨x, y, hxy⟩ := hR.exists_pair_ne in hxy (polynomial.ext (λ n, hR' _ _)) } + nontriviality R, + intro hR, + obtain ⟨p, hp⟩ := hR.mul_inv_cancel X_ne_zero, + have hp0 : p ≠ 0, + { rintro rfl, + rw [mul_zero] at hp, + exact zero_ne_one hp }, + have := degree_lt_degree_mul_X hp0, + rw [←X_mul, congr_arg degree hp, degree_one, nat.with_bot.lt_zero_iff, degree_eq_bot] at this, + exact hp0 this, end /-- The only constant in a maximal ideal over a field is `0`. -/ From 1a862493f1cd226be1f2e26d46e4a53a57c26436 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 4 May 2022 07:51:01 +0000 Subject: [PATCH 091/153] feat(measure_theory/function/l1_space): add `integrable_smul_measure` (#13922) * add `integrable_smul_measure`, an `iff` version of `integrable.smul_measure`; * add `integrable_average`, an `iff` version of `integrable.to_average`. --- src/measure_theory/function/l1_space.lean | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/src/measure_theory/function/l1_space.lean b/src/measure_theory/function/l1_space.lean index 13dc10897967d..eb5066f28b773 100644 --- a/src/measure_theory/function/l1_space.lean +++ b/src/measure_theory/function/l1_space.lean @@ -505,6 +505,11 @@ lemma integrable.smul_measure {f : α → β} (h : integrable f μ) {c : ℝ≥0 integrable f (c • μ) := by { rw ← mem_ℒp_one_iff_integrable at h ⊢, exact h.smul_measure hc, } +lemma integrable_smul_measure {f : α → β} {c : ℝ≥0∞} (h₁ : c ≠ 0) (h₂ : c ≠ ∞) : + integrable f (c • μ) ↔ integrable f μ := +⟨λ h, by simpa only [smul_smul, ennreal.inv_mul_cancel h₁ h₂, one_smul] + using h.smul_measure (ennreal.inv_ne_top.2 h₁), λ h, h.smul_measure h₂⟩ + lemma integrable.to_average {f : α → β} (h : integrable f μ) : integrable f ((μ univ)⁻¹ • μ) := begin @@ -513,6 +518,12 @@ begin { apply h.smul_measure, simpa } end +lemma integrable_average [is_finite_measure μ] {f : α → β} : + integrable f ((μ univ)⁻¹ • μ) ↔ integrable f μ := +(eq_or_ne μ 0).by_cases (λ h, by simp [h]) $ λ h, + integrable_smul_measure (ennreal.inv_ne_zero.2 $ measure_ne_top _ _) + (ennreal.inv_ne_top.2 $ mt measure.measure_univ_eq_zero.1 h) + lemma integrable_map_measure {f : α → δ} {g : δ → β} (hg : ae_strongly_measurable g (measure.map f μ)) (hf : ae_measurable f μ) : integrable g (measure.map f μ) ↔ integrable (g ∘ f) μ := From e24f7f7ae84b6b7987073bd99ff4ed60ef469124 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 4 May 2022 11:10:40 +0000 Subject: [PATCH 092/153] =?UTF-8?q?move(set=5Ftheory/ordinal/{arithmetic?= =?UTF-8?q?=20=E2=86=92=20fixed=5Fpoints}):=20Move=20`nfp`=20(#13315)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit That way, it belong with the other functions about fixed points. --- src/set_theory/ordinal/arithmetic.lean | 298 ----------------------- src/set_theory/ordinal/fixed_point.lean | 306 +++++++++++++++++++++++- src/set_theory/ordinal/principal.lean | 4 +- 3 files changed, 298 insertions(+), 310 deletions(-) diff --git a/src/set_theory/ordinal/arithmetic.lean b/src/set_theory/ordinal/arithmetic.lean index cb182df7f9126..46c56dc9097cd 100644 --- a/src/set_theory/ordinal/arithmetic.lean +++ b/src/set_theory/ordinal/arithmetic.lean @@ -40,8 +40,6 @@ Some properties of the operations are also used to discuss general tools on ordi and order-continuous, i.e., the image `f o` of a limit ordinal `o` is the sup of `f a` for `a < o`. * `enum_ord`: enumerates an unbounded set of ordinals by the ordinals themselves. -* `nfp f a`: the next fixed point of a function `f` on ordinals, above `a`. It behaves well - for normal functions. * `CNF b o` is the Cantor normal form of the ordinal `o` in base `b`. * `sup`, `lsub`: the supremum / least strict upper bound of an indexed family of ordinals in `Type u`, as an ordinal in `Type u`. @@ -2420,300 +2418,4 @@ begin rw [nat.cast_zero, opow_zero] } end -/-! ### Fixed points of normal functions -/ - -section -variable {f : ordinal.{u} → ordinal.{u}} - -/-- The next fixed point function, the least fixed point of the normal function `f` above `a`. -/ -def nfp (f : ordinal → ordinal) (a : ordinal) := -sup (λ n : ℕ, f^[n] a) - -theorem nfp_le_iff {f a b} : nfp f a ≤ b ↔ ∀ n, f^[n] a ≤ b := -sup_le_iff - -theorem nfp_le {f a b} : (∀ n, f^[n] a ≤ b) → nfp f a ≤ b := -sup_le - -theorem iterate_le_nfp (f a n) : f^[n] a ≤ nfp f a := -le_sup _ n - -theorem le_nfp_self (f a) : a ≤ nfp f a := -iterate_le_nfp f a 0 - -theorem lt_nfp {f : ordinal → ordinal} {a b} : a < nfp f b ↔ ∃ n, a < (f^[n]) b := -lt_sup - -protected theorem is_normal.lt_nfp {f} (H : is_normal f) {a b} : f b < nfp f a ↔ b < nfp f a := -lt_nfp.trans $ iff.trans - (by exact - ⟨λ ⟨n, h⟩, ⟨n, lt_of_le_of_lt (H.self_le _) h⟩, - λ ⟨n, h⟩, ⟨n+1, by rw iterate_succ'; exact H.lt_iff.2 h⟩⟩) - lt_sup.symm - -protected theorem is_normal.nfp_le (H : is_normal f) {a b} : nfp f a ≤ f b ↔ nfp f a ≤ b := -le_iff_le_iff_lt_iff_lt.2 H.lt_nfp - -theorem is_normal.nfp_le_fp (H : is_normal f) {a b} (ab : a ≤ b) (h : f b ≤ b) : nfp f a ≤ b := -nfp_le $ λ i, begin - induction i with i IH generalizing a, {exact ab}, - exact IH (le_trans (H.le_iff.2 ab) h), -end - -theorem is_normal.nfp_fp (H : is_normal f) (a) : f (nfp f a) = nfp f a := -begin - refine le_antisymm _ (H.self_le _), - cases le_or_lt (f a) a with aa aa, - { rwa le_antisymm (H.nfp_le_fp le_rfl aa) (le_nfp_self _ _) }, - rcases zero_or_succ_or_limit (nfp f a) with e|⟨b, e⟩|l, - { refine @le_trans _ _ _ (f a) _ (H.le_iff.2 _) (iterate_le_nfp f a 1), - simp only [e, ordinal.zero_le] }, - { have : f b < nfp f a := H.lt_nfp.2 (by simp only [e, lt_succ_self]), - rw [e, lt_succ] at this, - have ab : a ≤ b, - { rw [← lt_succ, ← e], - exact lt_of_lt_of_le aa (iterate_le_nfp f a 1) }, - refine le_trans (H.le_iff.2 (H.nfp_le_fp ab this)) - (le_trans this (le_of_lt _)), - simp only [e, lt_succ_self] }, - { exact (H.2 _ l _).2 (λ b h, le_of_lt (H.lt_nfp.2 h)) } -end - -theorem is_normal.le_nfp (H : is_normal f) {a b} : f b ≤ nfp f a ↔ b ≤ nfp f a := -⟨le_trans (H.self_le _), λ h, by simpa only [H.nfp_fp] using H.le_iff.2 h⟩ - -theorem nfp_eq_self {a} (h : f a = a) : nfp f a = a := -(nfp_le $ λ i, by rw [iterate_fixed h]).antisymm (le_nfp_self f a) - -protected lemma monotone.nfp (hf : monotone f) : monotone (nfp f) := -λ a b h, nfp_le (λ n, (hf.iterate n h).trans (le_sup _ n)) - -/-- Fixed point lemma for normal functions: the fixed points of a normal function are unbounded. -/ -theorem is_normal.nfp_unbounded {f} (H : is_normal f) : unbounded (<) (fixed_points f) := -λ a, ⟨_, H.nfp_fp a, not_lt_of_ge (le_nfp_self f a)⟩ - -/-- The derivative of a normal function `f` is the sequence of fixed points of `f`. -/ -def deriv (f : ordinal → ordinal) (o : ordinal) : ordinal := -limit_rec_on o (nfp f 0) - (λ a IH, nfp f (succ IH)) - (λ a l, bsup.{u u} a) - -@[simp] theorem deriv_zero (f) : deriv f 0 = nfp f 0 := limit_rec_on_zero _ _ _ - -@[simp] theorem deriv_succ (f o) : deriv f (succ o) = nfp f (succ (deriv f o)) := -limit_rec_on_succ _ _ _ _ - -theorem deriv_limit (f) {o} : is_limit o → deriv f o = bsup.{u u} o (λ a _, deriv f a) := -limit_rec_on_limit _ _ _ _ - -theorem deriv_is_normal (f) : is_normal (deriv f) := -⟨λ o, by rw [deriv_succ, ← succ_le]; apply le_nfp_self, - λ o l a, by rw [deriv_limit _ l, bsup_le_iff]⟩ - -theorem is_normal.deriv_fp (H : is_normal f) (o) : f (deriv.{u} f o) = deriv f o := -begin - apply limit_rec_on o, - { rw [deriv_zero, H.nfp_fp] }, - { intros o ih, rw [deriv_succ, H.nfp_fp] }, - intros o l IH, - rw [deriv_limit _ l, is_normal.bsup.{u u u} H _ l.1], - refine eq_of_forall_ge_iff (λ c, _), - simp only [bsup_le_iff, IH] {contextual:=tt} -end - -theorem is_normal.le_iff_deriv (H : is_normal f) {a} : f a ≤ a ↔ ∃ o, deriv f o = a := -⟨λ ha, begin - suffices : ∀ o (_ : a ≤ deriv f o), ∃ o, deriv f o = a, - from this a ((deriv_is_normal _).self_le _), - refine λ o, limit_rec_on o (λ h₁, ⟨0, le_antisymm _ h₁⟩) (λ o IH h₁, _) (λ o l IH h₁, _), - { rw deriv_zero, - exact H.nfp_le_fp (ordinal.zero_le _) ha }, - { cases le_or_lt a (deriv f o), {exact IH h}, - refine ⟨succ o, le_antisymm _ h₁⟩, - rw deriv_succ, - exact H.nfp_le_fp (succ_le.2 h) ha }, - { cases eq_or_lt_of_le h₁, {exact ⟨_, h.symm⟩}, - rw [deriv_limit _ l, ← not_le, bsup_le_iff, not_ball] at h, - exact let ⟨o', h, hl⟩ := h in IH o' h (le_of_not_le hl) } -end, λ ⟨o, e⟩, e ▸ le_of_eq (H.deriv_fp _)⟩ - -theorem is_normal.apply_eq_self_iff_deriv (H : is_normal f) {a} : f a = a ↔ ∃ o, deriv f o = a := -by rw [←H.le_iff_deriv, H.le_iff_eq] - -/-- `deriv f` is the fixed point enumerator of `f`. -/ -theorem deriv_eq_enum_fp {f} (H : is_normal f) : deriv f = enum_ord (fixed_points f) := -begin - rw [←eq_enum_ord _ H.nfp_unbounded, range_eq_iff], - exact ⟨(deriv_is_normal f).strict_mono, H.deriv_fp, λ _, H.apply_eq_self_iff_deriv.1⟩ -end - -theorem deriv_eq_id_of_nfp_eq_id {f : ordinal → ordinal} (h : nfp f = id) : deriv f = id := -(is_normal.eq_iff_zero_and_succ (deriv_is_normal _) is_normal.refl).2 (by simp [h, succ_inj]) - -end - -/-! ### Fixed points of addition -/ - -@[simp] theorem nfp_add_zero (a) : nfp ((+) a) 0 = a * omega := -begin - unfold nfp, - rw ←sup_mul_nat, - congr, funext, - induction n with n hn, - { rw [nat.cast_zero, mul_zero, iterate_zero_apply] }, - { nth_rewrite 1 nat.succ_eq_one_add, - rw [nat.cast_add, nat.cast_one, mul_one_add, iterate_succ_apply', hn] } -end - -theorem nfp_add_eq_mul_omega {a b} (hba : b ≤ a * omega) : - nfp ((+) a) b = a * omega := -begin - apply le_antisymm ((add_is_normal a).nfp_le_fp hba _), - { rw ←nfp_add_zero, - exact monotone.nfp (add_is_normal a).monotone (ordinal.zero_le b) }, - { rw [←mul_one_add, one_add_omega] } -end - -theorem add_eq_right_iff_mul_omega_le {a b : ordinal} : a + b = b ↔ a * omega ≤ b := -begin - refine ⟨λ h, _, λ h, _⟩, - { rw [←nfp_add_zero a, ←deriv_zero], - cases (add_is_normal a).apply_eq_self_iff_deriv.1 h with c hc, - rw ←hc, - exact (deriv_is_normal _).monotone (ordinal.zero_le _) }, - { have := ordinal.add_sub_cancel_of_le h, - nth_rewrite 0 ←this, - rwa [←add_assoc, ←mul_one_add, one_add_omega] } -end - -theorem add_le_right_iff_mul_omega_le {a b : ordinal} : a + b ≤ b ↔ a * omega ≤ b := -by { rw ←add_eq_right_iff_mul_omega_le, exact (add_is_normal a).le_iff_eq } - -theorem deriv_add_eq_mul_omega_add (a b : ordinal.{u}) : deriv ((+) a) b = a * omega + b := -begin - revert b, - rw [←funext_iff, is_normal.eq_iff_zero_and_succ (deriv_is_normal _) (add_is_normal _)], - refine ⟨_, λ a h, _⟩, - { rw [deriv_zero, add_zero], - exact nfp_add_zero a }, - { rw [deriv_succ, h, add_succ], - exact nfp_eq_self (add_eq_right_iff_mul_omega_le.2 ((le_add_right _ _).trans - (lt_succ_self _).le)) } -end - -/-! ### Fixed points of multiplication -/ - -@[simp] theorem nfp_mul_one {a : ordinal} (ha : 0 < a) : nfp ((*) a) 1 = a ^ omega := -begin - unfold nfp, - rw ←sup_opow_nat, - { congr, - funext, - induction n with n hn, - { rw [nat.cast_zero, opow_zero, iterate_zero_apply] }, - nth_rewrite 1 nat.succ_eq_one_add, - rw [nat.cast_add, nat.cast_one, opow_add, opow_one, iterate_succ_apply', hn] }, - { exact ha } -end - -@[simp] theorem nfp_mul_zero (a : ordinal) : nfp ((*) a) 0 = 0 := -begin - rw [←ordinal.le_zero, nfp_le_iff], - intro n, - induction n with n hn, { refl }, - rwa [iterate_succ_apply, mul_zero] -end - -@[simp] theorem nfp_zero_mul : nfp ((*) 0) = id := -begin - refine funext (λ a, (nfp_le (λ n, _)).antisymm (le_sup (λ n, ((*) 0)^[n] a) 0)), - induction n with n hn, { refl }, - rw function.iterate_succ', - change 0 * _ ≤ a, - rw zero_mul, - exact ordinal.zero_le a -end - -@[simp] theorem deriv_mul_zero : deriv ((*) 0) = id := -deriv_eq_id_of_nfp_eq_id nfp_zero_mul - -theorem nfp_mul_eq_opow_omega {a b : ordinal} (hb : 0 < b) (hba : b ≤ a ^ omega) : - nfp ((*) a) b = a ^ omega.{u} := -begin - cases eq_zero_or_pos a with ha ha, - { rw [ha, zero_opow omega_ne_zero] at *, - rw [ordinal.le_zero.1 hba, nfp_zero_mul], - refl }, - apply le_antisymm, - { apply (mul_is_normal ha).nfp_le_fp hba, - rw [←opow_one_add, one_add_omega] }, - rw ←nfp_mul_one ha, - exact monotone.nfp (mul_is_normal ha).monotone (one_le_iff_pos.2 hb) -end - -theorem eq_zero_or_opow_omega_le_of_mul_eq_right {a b : ordinal} (hab : a * b = b) : - b = 0 ∨ a ^ omega.{u} ≤ b := -begin - cases eq_zero_or_pos a with ha ha, - { rw [ha, zero_opow omega_ne_zero], - exact or.inr (ordinal.zero_le b) }, - rw or_iff_not_imp_left, - intro hb, - change b ≠ 0 at hb, - rw ←nfp_mul_one ha, - rw ←one_le_iff_ne_zero at hb, - exact (mul_is_normal ha).nfp_le_fp hb (le_of_eq hab) -end - -theorem mul_eq_right_iff_opow_omega_dvd {a b : ordinal} : a * b = b ↔ a ^ omega ∣ b := -begin - cases eq_zero_or_pos a with ha ha, - { rw [ha, zero_mul, zero_opow omega_ne_zero, zero_dvd], - exact eq_comm }, - refine ⟨λ hab, _, λ h, _⟩, - { rw dvd_iff_mod_eq_zero, - rw [←div_add_mod b (a ^ omega), mul_add, ←mul_assoc, ←opow_one_add, one_add_omega, - add_left_cancel] at hab, - cases eq_zero_or_opow_omega_le_of_mul_eq_right hab with hab hab, - { exact hab }, - refine (not_lt_of_le hab (mod_lt b (opow_ne_zero omega _))).elim, - rwa ←ordinal.pos_iff_ne_zero }, - cases h with c hc, - rw [hc, ←mul_assoc, ←opow_one_add, one_add_omega] -end - -theorem mul_le_right_iff_opow_omega_dvd {a b : ordinal} (ha : 0 < a) : a * b ≤ b ↔ a ^ omega ∣ b := -by { rw ←mul_eq_right_iff_opow_omega_dvd, exact (mul_is_normal ha).le_iff_eq } - -theorem nfp_mul_opow_omega_add {a c : ordinal} (b) (ha : 0 < a) (hc : 0 < c) (hca : c ≤ a ^ omega) : - nfp ((*) a) (a ^ omega * b + c) = a ^ omega.{u} * b.succ := -begin - apply le_antisymm, - { apply is_normal.nfp_le_fp (mul_is_normal ha), - { rw mul_succ, - apply add_le_add_left hca }, - { rw [←mul_assoc, ←opow_one_add, one_add_omega] } }, - { cases mul_eq_right_iff_opow_omega_dvd.1 ((mul_is_normal ha).nfp_fp (a ^ omega * b + c)) - with d hd, - rw hd, - apply mul_le_mul_left', - have := le_nfp_self (has_mul.mul a) (a ^ omega * b + c), - rw hd at this, - have := (add_lt_add_left hc (a ^ omega * b)).trans_le this, - rw [add_zero, mul_lt_mul_iff_left (opow_pos omega ha)] at this, - rwa succ_le } -end - -theorem deriv_mul_eq_opow_omega_mul {a : ordinal.{u}} (ha : 0 < a) (b) : - deriv ((*) a) b = a ^ omega * b := -begin - revert b, - rw [←funext_iff, - is_normal.eq_iff_zero_and_succ (deriv_is_normal _) (mul_is_normal (opow_pos omega ha))], - refine ⟨_, λ c h, _⟩, - { rw [deriv_zero, nfp_mul_zero, mul_zero] }, - { rw [deriv_succ, h], - exact nfp_mul_opow_omega_add c ha zero_lt_one (one_le_iff_pos.2 (opow_pos _ ha)) }, -end - end ordinal diff --git a/src/set_theory/ordinal/fixed_point.lean b/src/set_theory/ordinal/fixed_point.lean index 13359187aec24..d5e81eb1092fb 100644 --- a/src/set_theory/ordinal/fixed_point.lean +++ b/src/set_theory/ordinal/fixed_point.lean @@ -1,7 +1,7 @@ /- -Copyright (c) 2022 Violeta Hernández Palacios. All rights reserved. +Copyright (c) 2018 Violeta Hernández Palacios, Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Violeta Hernández Palacios +Authors: Violeta Hernández Palacios, Mario Carneiro -/ import set_theory.ordinal.arithmetic @@ -22,6 +22,7 @@ Moreover, we prove some lemmas about the fixed points of specific normal functio * `fp_family_unbounded`, `fp_bfamily_unbounded`, `fp_unbounded`: the (common) fixed points of a (family of) normal function(s) are unbounded in the ordinals. * `deriv_add_eq_mul_omega_add`: a characterization of the derivative of addition. +* `deriv_mul_eq_opow_omega_mul`: a characterization of the derivative of multiplication. -/ noncomputable theory @@ -39,7 +40,7 @@ variables {ι : Type u} {f : ι → ordinal.{max u v} → ordinal.{max u v}} /-- The next common fixed point, at least `a`, for a family of normal functions. -`ordinal.nfp_family_fp` shows this is a fixed point, `ordinal.self_le_nfp_family` shows it's at +`ordinal.nfp_family_fp` shows this is a fixed point, `ordinal.le_nfp_family` shows it's at least `a`, and `ordinal.nfp_family_le_fp` shows this is the least ordinal with these properties. -/ def nfp_family (f : ι → ordinal → ordinal) (a) : ordinal := sup (list.foldr f a) @@ -52,7 +53,7 @@ theorem foldr_le_nfp_family (f : ι → ordinal → ordinal) (a l) : list.foldr f a l ≤ nfp_family f a := le_sup _ _ -theorem self_le_nfp_family (f : ι → ordinal → ordinal) (a) : a ≤ nfp_family f a := +theorem le_nfp_family (f : ι → ordinal → ordinal) (a) : a ≤ nfp_family f a := le_sup _ [] theorem lt_nfp_family {a b} : a < nfp_family f b ↔ ∃ l, a < list.foldr f b l := @@ -112,7 +113,7 @@ end theorem nfp_family_eq_self {f : ι → ordinal → ordinal} {a} (h : ∀ i, f i a = a) : nfp_family f a = a := -le_antisymm (sup_le (λ l, (by rw list.foldr_fixed' h l))) (self_le_nfp_family f a) +le_antisymm (sup_le (λ l, (by rw list.foldr_fixed' h l))) (le_nfp_family f a) /-- A generalization of the fixed point lemma for normal functions: any family of normal functions has an unbounded set of common fixed points. -/ @@ -121,7 +122,7 @@ theorem fp_family_unbounded (H : ∀ i, is_normal (f i)) : λ a, ⟨_, λ s ⟨i, hi⟩, begin rw ←hi, exact nfp_family_fp (H i) a -end, (self_le_nfp_family f a).not_lt⟩ +end, (le_nfp_family f a).not_lt⟩ /-- The derivative of a family of normal functions is the sequence of their common fixed points. -/ def deriv_family (f : ι → ordinal → ordinal) (o : ordinal) : ordinal := @@ -142,7 +143,7 @@ theorem deriv_family_limit (f : ι → ordinal → ordinal) {o} : is_limit o → limit_rec_on_limit _ _ _ _ theorem deriv_family_is_normal (f : ι → ordinal → ordinal) : is_normal (deriv_family f) := -⟨λ o, by rw [deriv_family_succ, ← succ_le]; apply self_le_nfp_family, +⟨λ o, by rw [deriv_family_succ, ← succ_le]; apply le_nfp_family, λ o l a, by rw [deriv_family_limit _ l, bsup_le_iff]⟩ theorem deriv_family_fp {i} (H : is_normal (f i)) (o : ordinal.{max u v}) : @@ -213,7 +214,7 @@ theorem foldr_le_nfp_bfamily {o : ordinal} (f : Π b < o, ordinal → ordinal) ( list.foldr (family_of_bfamily o f) a l ≤ nfp_bfamily o f a := le_sup _ _ -theorem self_le_nfp_bfamily {o : ordinal} (f : Π b < o, ordinal → ordinal) (a) : +theorem le_nfp_bfamily {o : ordinal} (f : Π b < o, ordinal → ordinal) (a) : a ≤ nfp_bfamily o f a := le_sup _ [] @@ -255,7 +256,7 @@ theorem nfp_bfamily_fp {i hi} (H : is_normal (f i hi)) (a) : f i hi (nfp_bfamily o f a) = nfp_bfamily o f a := by { rw ←family_of_bfamily_enum o f, apply nfp_family_fp, rw family_of_bfamily_enum, exact H } -theorem le_nfp_bfamily (ho : o ≠ 0) (H : ∀ i hi, is_normal (f i hi)) {a b} : +theorem apply_le_nfp_bfamily (ho : o ≠ 0) (H : ∀ i hi, is_normal (f i hi)) {a b} : (∀ i hi, f i hi b ≤ nfp_bfamily o f a) ↔ b ≤ nfp_bfamily o f a := begin refine ⟨λ h, _, λ h i hi, _⟩, @@ -273,7 +274,7 @@ nfp_family_eq_self (λ _, h _ _) theorem fp_bfamily_unbounded (H : ∀ i hi, is_normal (f i hi)) : (⋂ i hi, function.fixed_points (f i hi)).unbounded (<) := λ a, ⟨_, by { rw set.mem_Inter₂, exact λ i hi, nfp_bfamily_fp (H i hi) _ }, - (self_le_nfp_bfamily f a).not_lt⟩ + (le_nfp_bfamily f a).not_lt⟩ /-- The derivative of a family of normal functions is the sequence of their common fixed points. -/ def deriv_bfamily (o : ordinal) (f : Π b < o, ordinal → ordinal) : ordinal → ordinal := @@ -323,4 +324,289 @@ begin end end + +/-! ### Fixed points of a single function -/ + +section +variable {f : ordinal.{u} → ordinal.{u}} + +/-- The next fixed point function, the least fixed point of the normal function `f`, at least `a`. +-/ +def nfp (f : ordinal → ordinal) : ordinal → ordinal := +nfp_family (λ _ : unit, f) + +theorem nfp_eq_nfp_family (f : ordinal → ordinal) : nfp f = nfp_family (λ _ : unit, f) := +rfl + +@[simp] theorem sup_iterate_eq_nfp (f : ordinal.{u} → ordinal.{u}) : + (λ a, sup (λ n : ℕ, f^[n] a)) = nfp f := +begin + refine funext (λ a, le_antisymm _ (sup_le (λ l, _))), + { rw sup_le_iff, + intro n, + rw [←list.length_repeat unit.star n, ←list.foldr_const f a], + apply le_sup }, + { rw list.foldr_const f a l, + exact le_sup _ _ }, +end + +theorem iterate_le_nfp (f a n) : f^[n] a ≤ nfp f a := +by { rw ←sup_iterate_eq_nfp, exact le_sup _ n } + +theorem le_nfp (f a) : a ≤ nfp f a := +iterate_le_nfp f a 0 + +theorem lt_nfp {a b} : a < nfp f b ↔ ∃ n, a < (f^[n]) b := +by { rw ←sup_iterate_eq_nfp, exact lt_sup } + +theorem nfp_le_iff {a b} : nfp f a ≤ b ↔ ∀ n, (f^[n]) a ≤ b := +by { rw ←sup_iterate_eq_nfp, exact sup_le_iff } + +theorem nfp_le {a b} : (∀ n, (f^[n]) a ≤ b) → nfp f a ≤ b := +nfp_le_iff.2 + +@[simp] theorem nfp_id : nfp id = id := +funext (λ a, begin + simp_rw [←sup_iterate_eq_nfp, iterate_id], + exact sup_const a +end) + +theorem nfp_monotone (hf : monotone f) : monotone (nfp f) := +nfp_family_monotone (λ i, hf) + +theorem is_normal.apply_lt_nfp {f} (H : is_normal f) {a b} : + f b < nfp f a ↔ b < nfp f a := +begin + unfold nfp, + rw ←@apply_lt_nfp_family_iff unit (λ _, f) _ (λ _, H) a b, + exact ⟨λ h _, h, λ h, h unit.star⟩ +end + +theorem is_normal.nfp_le_apply {f} (H : is_normal f) {a b} : nfp f a ≤ f b ↔ nfp f a ≤ b := +le_iff_le_iff_lt_iff_lt.2 H.apply_lt_nfp + +theorem nfp_le_fp {f} (H : monotone f) {a b} (ab : a ≤ b) (h : f b ≤ b) : nfp f a ≤ b := +nfp_family_le_fp (λ _, H) ab (λ _, h) + +theorem is_normal.nfp_fp {f} (H : is_normal f) : ∀ a, f (nfp f a) = nfp f a := +@nfp_family_fp unit (λ _, f) unit.star H + +theorem is_normal.apply_le_nfp {f} (H : is_normal f) {a b} : + f b ≤ nfp f a ↔ b ≤ nfp f a := +⟨le_trans (H.self_le _), λ h, by simpa only [H.nfp_fp] using H.le_iff.2 h⟩ + +theorem nfp_eq_self {f : ordinal → ordinal} {a} (h : f a = a) : nfp f a = a := +nfp_family_eq_self (λ _, h) + +/-- The fixed point lemma for normal functions: any normal function has an unbounded set of +fixed points. -/ +theorem fp_unbounded (H : is_normal f) : (function.fixed_points f).unbounded (<) := +by { convert fp_family_unbounded (λ _ : unit, H), exact (set.Inter_const _).symm } + +/-- The derivative of a normal function `f` is the sequence of fixed points of `f`. -/ +def deriv (f : ordinal → ordinal) : ordinal → ordinal := +deriv_family (λ _ : unit, f) + +theorem deriv_eq_deriv_family (f : ordinal → ordinal) : deriv f = deriv_family (λ _ : unit, f) := +rfl + +@[simp] theorem deriv_zero (f) : deriv f 0 = nfp f 0 := +deriv_family_zero _ + +@[simp] theorem deriv_succ (f o) : deriv f (succ o) = nfp f (succ (deriv f o)) := +deriv_family_succ _ _ + +theorem deriv_limit (f) {o} : is_limit o → deriv f o = bsup.{u 0} o (λ a _, deriv f a) := +deriv_family_limit _ + +theorem deriv_is_normal (f) : is_normal (deriv f) := +deriv_family_is_normal _ + +theorem deriv_id_of_nfp_id {f : ordinal → ordinal} (h : nfp f = id) : deriv f = id := +((deriv_is_normal _).eq_iff_zero_and_succ is_normal.refl).2 (by simp [h, succ_inj]) + +theorem is_normal.deriv_fp {f} (H : is_normal f) : ∀ o, f (deriv f o) = deriv f o := +@deriv_family_fp unit (λ _, f) unit.star H + +theorem is_normal.le_iff_deriv {f} (H : is_normal f) {a} : f a ≤ a ↔ ∃ o, deriv f o = a := +begin + unfold deriv, + rw ←le_iff_deriv_family (λ _ : unit, H), + exact ⟨λ h _, h, λ h, h unit.star⟩ +end + +theorem is_normal.fp_iff_deriv {f} (H : is_normal f) {a} : f a = a ↔ ∃ o, deriv f o = a := +by rw [←H.le_iff_eq, H.le_iff_deriv] + +theorem deriv_eq_enum_ord (H : is_normal f) : deriv f = enum_ord (function.fixed_points f) := +by { convert deriv_family_eq_enum_ord (λ _ : unit, H), exact (set.Inter_const _).symm } + +theorem deriv_eq_id_of_nfp_eq_id {f : ordinal → ordinal} (h : nfp f = id) : deriv f = id := +(is_normal.eq_iff_zero_and_succ (deriv_is_normal _) is_normal.refl).2 (by simp [h, succ_inj]) + +end + +/-! ### Fixed points of addition -/ + +@[simp] theorem nfp_add_zero (a) : nfp ((+) a) 0 = a * omega := +begin + simp_rw [←sup_iterate_eq_nfp, ←sup_mul_nat], + congr, funext, + induction n with n hn, + { rw [nat.cast_zero, mul_zero, iterate_zero_apply] }, + { nth_rewrite 1 nat.succ_eq_one_add, + rw [nat.cast_add, nat.cast_one, mul_one_add, iterate_succ_apply', hn] } +end + +theorem nfp_add_eq_mul_omega {a b} (hba : b ≤ a * omega) : + nfp ((+) a) b = a * omega := +begin + apply le_antisymm (nfp_le_fp (add_is_normal a).monotone hba _), + { rw ←nfp_add_zero, + exact nfp_monotone (add_is_normal a).monotone (ordinal.zero_le b) }, + { rw [←mul_one_add, one_add_omega] } +end + +theorem add_eq_right_iff_mul_omega_le {a b : ordinal} : a + b = b ↔ a * omega ≤ b := +begin + refine ⟨λ h, _, λ h, _⟩, + { rw [←nfp_add_zero a, ←deriv_zero], + cases (add_is_normal a).fp_iff_deriv.1 h with c hc, + rw ←hc, + exact (deriv_is_normal _).monotone (ordinal.zero_le _) }, + { have := ordinal.add_sub_cancel_of_le h, + nth_rewrite 0 ←this, + rwa [←add_assoc, ←mul_one_add, one_add_omega] } +end + +theorem add_le_right_iff_mul_omega_le {a b : ordinal} : a + b ≤ b ↔ a * omega ≤ b := +by { rw ←add_eq_right_iff_mul_omega_le, exact (add_is_normal a).le_iff_eq } + +theorem deriv_add_eq_mul_omega_add (a b : ordinal.{u}) : deriv ((+) a) b = a * omega + b := +begin + revert b, + rw [←funext_iff, is_normal.eq_iff_zero_and_succ (deriv_is_normal _) (add_is_normal _)], + refine ⟨_, λ a h, _⟩, + { rw [deriv_zero, add_zero], + exact nfp_add_zero a }, + { rw [deriv_succ, h, add_succ], + exact nfp_eq_self (add_eq_right_iff_mul_omega_le.2 ((le_add_right _ _).trans + (lt_succ_self _).le)) } +end + +/-! ### Fixed points of multiplication -/ + +local infixr ^ := @pow ordinal ordinal ordinal.has_pow +@[simp] theorem nfp_mul_one {a : ordinal} (ha : 0 < a) : nfp ((*) a) 1 = a ^ omega := +begin + rw [←sup_iterate_eq_nfp, ←sup_opow_nat], + { dsimp, congr, funext, + induction n with n hn, + { rw [nat.cast_zero, opow_zero, iterate_zero_apply] }, + nth_rewrite 1 nat.succ_eq_one_add, + rw [nat.cast_add, nat.cast_one, opow_add, opow_one, iterate_succ_apply', hn] }, + { exact ha } +end + +@[simp] theorem nfp_mul_zero (a : ordinal) : nfp ((*) a) 0 = 0 := +begin + rw [←ordinal.le_zero, nfp_le_iff], + intro n, + induction n with n hn, { refl }, + rwa [iterate_succ_apply, mul_zero] +end + +@[simp] theorem nfp_zero_mul : nfp ((*) 0) = id := +begin + rw ←sup_iterate_eq_nfp, + refine funext (λ a, (sup_le (λ n, _)).antisymm (le_sup (λ n, ((*) 0)^[n] a) 0)), + induction n with n hn, { refl }, + rw function.iterate_succ', + change 0 * _ ≤ a, + rw zero_mul, + exact ordinal.zero_le a +end + +@[simp] theorem deriv_mul_zero : deriv ((*) 0) = id := +deriv_eq_id_of_nfp_eq_id nfp_zero_mul + +theorem nfp_mul_eq_opow_omega {a b : ordinal} (hb : 0 < b) (hba : b ≤ a ^ omega) : + nfp ((*) a) b = a ^ omega.{u} := +begin + cases eq_zero_or_pos a with ha ha, + { rw [ha, zero_opow omega_ne_zero] at *, + rw [ordinal.le_zero.1 hba, nfp_zero_mul], + refl }, + apply le_antisymm, + { apply nfp_le_fp (mul_is_normal ha).monotone hba, + rw [←opow_one_add, one_add_omega] }, + rw ←nfp_mul_one ha, + exact nfp_monotone (mul_is_normal ha).monotone (one_le_iff_pos.2 hb) +end + +theorem eq_zero_or_opow_omega_le_of_mul_eq_right {a b : ordinal} (hab : a * b = b) : + b = 0 ∨ a ^ omega.{u} ≤ b := +begin + cases eq_zero_or_pos a with ha ha, + { rw [ha, zero_opow omega_ne_zero], + exact or.inr (ordinal.zero_le b) }, + rw or_iff_not_imp_left, + intro hb, + change b ≠ 0 at hb, + rw ←nfp_mul_one ha, + rw ←one_le_iff_ne_zero at hb, + exact nfp_le_fp (mul_is_normal ha).monotone hb (le_of_eq hab) +end + +theorem mul_eq_right_iff_opow_omega_dvd {a b : ordinal} : a * b = b ↔ a ^ omega ∣ b := +begin + cases eq_zero_or_pos a with ha ha, + { rw [ha, zero_mul, zero_opow omega_ne_zero, zero_dvd], + exact eq_comm }, + refine ⟨λ hab, _, λ h, _⟩, + { rw dvd_iff_mod_eq_zero, + rw [←div_add_mod b (a ^ omega), mul_add, ←mul_assoc, ←opow_one_add, one_add_omega, + add_left_cancel] at hab, + cases eq_zero_or_opow_omega_le_of_mul_eq_right hab with hab hab, + { exact hab }, + refine (not_lt_of_le hab (mod_lt b (opow_ne_zero omega _))).elim, + rwa ←ordinal.pos_iff_ne_zero }, + cases h with c hc, + rw [hc, ←mul_assoc, ←opow_one_add, one_add_omega] +end + +theorem mul_le_right_iff_opow_omega_dvd {a b : ordinal} (ha : 0 < a) : a * b ≤ b ↔ a ^ omega ∣ b := +by { rw ←mul_eq_right_iff_opow_omega_dvd, exact (mul_is_normal ha).le_iff_eq } + +theorem nfp_mul_opow_omega_add {a c : ordinal} (b) (ha : 0 < a) (hc : 0 < c) (hca : c ≤ a ^ omega) : + nfp ((*) a) (a ^ omega * b + c) = a ^ omega.{u} * b.succ := +begin + apply le_antisymm, + { apply nfp_le_fp (mul_is_normal ha).monotone, + { rw mul_succ, + apply add_le_add_left hca }, + { rw [←mul_assoc, ←opow_one_add, one_add_omega] } }, + { cases mul_eq_right_iff_opow_omega_dvd.1 ((mul_is_normal ha).nfp_fp (a ^ omega * b + c)) + with d hd, + rw hd, + apply mul_le_mul_left', + have := le_nfp (has_mul.mul a) (a ^ omega * b + c), + rw hd at this, + have := (add_lt_add_left hc (a ^ omega * b)).trans_le this, + rw [add_zero, mul_lt_mul_iff_left (opow_pos omega ha)] at this, + rwa succ_le } +end + +theorem deriv_mul_eq_opow_omega_mul {a : ordinal.{u}} (ha : 0 < a) (b) : + deriv ((*) a) b = a ^ omega * b := +begin + revert b, + rw [←funext_iff, + is_normal.eq_iff_zero_and_succ (deriv_is_normal _) (mul_is_normal (opow_pos omega ha))], + refine ⟨_, λ c h, _⟩, + { rw [deriv_zero, nfp_mul_zero, mul_zero] }, + { rw [deriv_succ, h], + exact nfp_mul_opow_omega_add c ha zero_lt_one (one_le_iff_pos.2 (opow_pos _ ha)) }, +end + end ordinal diff --git a/src/set_theory/ordinal/principal.lean b/src/set_theory/ordinal/principal.lean index 578c9eec1786d..456ee103d9fd9 100644 --- a/src/set_theory/ordinal/principal.lean +++ b/src/set_theory/ordinal/principal.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Violeta Hernández Palacios -/ -import set_theory.ordinal.arithmetic +import set_theory.ordinal.fixed_point /-! ### Principal ordinals @@ -109,7 +109,7 @@ end theorem unbounded_principal (op : ordinal → ordinal → ordinal) : set.unbounded (<) {o | principal op o} := -λ o, ⟨_, principal_nfp_blsub₂ op o, (le_nfp_self _ o).not_lt⟩ +λ o, ⟨_, principal_nfp_blsub₂ op o, (le_nfp _ o).not_lt⟩ /-! #### Additive principal ordinals -/ From 6c7b880ed2cd2c48fe92576b33a48a29150d565b Mon Sep 17 00:00:00 2001 From: pbazin Date: Wed, 4 May 2022 11:10:42 +0000 Subject: [PATCH 093/153] feat(algebra/module/torsion): torsion ideal, decomposition lemma (#13414) Defines the torsion ideal of an element in a module, and also shows a decomposition lemma for torsion modules. Co-authored-by: pbazin <75327486+pbazin@users.noreply.github.com> --- src/algebra/module/torsion.lean | 276 +++++++++++++++++++++++++--- src/ring_theory/coprime/lemmas.lean | 12 +- 2 files changed, 259 insertions(+), 29 deletions(-) diff --git a/src/algebra/module/torsion.lean b/src/algebra/module/torsion.lean index e3883d9291d2c..aac753fdcdd60 100644 --- a/src/algebra/module/torsion.lean +++ b/src/algebra/module/torsion.lean @@ -7,6 +7,9 @@ import algebra.module import linear_algebra.quotient import ring_theory.ideal.quotient import ring_theory.non_zero_divisors +import algebra.direct_sum.module +import group_theory.torsion +import linear_algebra.isomorphisms import group_theory.torsion /-! @@ -14,21 +17,30 @@ import group_theory.torsion ## Main definitions +* `torsion_of R M x` : the torsion ideal of `x`, containing all `a` such that `a • x = 0`. * `submodule.torsion_by R M a` : the `a`-torsion submodule, containing all elements `x` of `M` such that `a • x = 0`. +* `submodule.torsion_by_set R M s` : the submodule containing all elements `x` of `M` such that + `a • x = 0` for all `a` in `s`. * `submodule.torsion' R M S` : the `S`-torsion submodule, containing all elements `x` of `M` such that `a • x = 0` for some `a` in `S`. * `submodule.torsion R M` : the torsion submoule, containing all elements `x` of `M` such that `a • x = 0` for some non-zero-divisor `a` in `R`. * `module.is_torsion_by R M a` : the property that defines a `a`-torsion module. Similarly, - `is_torsion'` and `is_torsion`. + `is_torsion_by_set`, `is_torsion'` and `is_torsion`. +* `module.is_torsion_by_set.module` : Creates a `R ⧸ I`-module from a `R`-module that + `is_torsion_by_set R _ I`. ## Main statements +* `quot_torsion_of_equiv_span_singleton` : isomorphism between the span of an element of `M` and + the quotient by its torsion ideal. * `torsion' R M S` and `torsion R M` are submodules. -* `torsion_by R M a` can be viewed as a `(R ⧸ R∙a)`-module. +* `torsion_by_set_eq_torsion_by_span` : torsion by a set is torsion by the ideal generated by it. * `submodule.torsion_by_is_torsion_by` : the `a`-torsion submodule is a `a`-torsion module. Similar lemmas for `torsion'` and `torsion`. +* `submodule.torsion_is_internal` : a `∏ i, p i`-torsion module is the internal direct sum of its + `p i`-torsion submodules when the `p i` are pairwise coprime. * `submodule.no_zero_smul_divisors_iff_torsion_bot` : a module over a domain has `no_zero_smul_divisors` (that is, there is no non-zero `a`, `x` such that `a • x = 0`) iff its torsion submodule is trivial. @@ -48,6 +60,27 @@ import group_theory.torsion Torsion, submodule, module, quotient -/ +section +variables (R M : Type*) [semiring R] [add_comm_monoid M] [module R M] +/--The torsion ideal of `x`, containing all `a` such that `a • x = 0`.-/ +@[simps] def torsion_of (x : M) : ideal R := (linear_map.to_span_singleton R M x).ker +variables {R M} +@[simp] lemma mem_torsion_of_iff (x : M) (a : R) : a ∈ torsion_of R M x ↔ a • x = 0 := iff.rfl +end + +section +variables (R M : Type*) [ring R] [add_comm_group M] [module R M] +/--The span of `x` in `M` is isomorphic to `R` quotiented by the torsion ideal of `x`.-/ +noncomputable def quot_torsion_of_equiv_span_singleton (x : M) : + (R ⧸ torsion_of R M x) ≃ₗ[R] (R ∙ x) := +(linear_map.to_span_singleton R M x).quot_ker_equiv_range.trans $ +linear_equiv.of_eq _ _ (linear_map.span_singleton_eq_range R M x).symm + +@[simp] lemma quot_torsion_of_equiv_span_singleton_apply_mk (x : M) (a : R) : + quot_torsion_of_equiv_span_singleton R M x (submodule.quotient.mk a) = + a • ⟨x, submodule.mem_span_singleton_self x⟩ := rfl +end + open_locale non_zero_divisors section defs @@ -60,6 +93,9 @@ namespace submodule `a • x = 0`. -/ @[simps] def torsion_by (a : R) : submodule R M := (distrib_mul_action.to_linear_map _ _ a).ker +/-- The submodule containing all elements `x` of `M` such that `a • x = 0` for all `a` in `s`. -/ +@[simps] def torsion_by_set (s : set R) : submodule R M := Inf (torsion_by R M '' s) + /-- The `S`-torsion submodule, containing all elements `x` of `M` such that `a • x = 0` for some `a` in `S`. -/ @[simps] def torsion' (S : Type*) @@ -82,6 +118,9 @@ namespace module /-- A `a`-torsion module is a module where every element is `a`-torsion. -/ @[reducible] def is_torsion_by (a : R) := ∀ ⦃x : M⦄, a • x = 0 +/-- A module where every element is `a`-torsion for all `a` in `s`. -/ +@[reducible] def is_torsion_by_set (s : set R) := ∀ ⦃x : M⦄ ⦃a : s⦄, (a : R) • x = 0 + /-- A `S`-torsion module is a module where every element is `a`-torsion for some `a` in `S`. -/ @[reducible] def is_torsion' (S : Type*) [has_scalar S M] := ∀ ⦃x : M⦄, ∃ a : S, a • x = 0 @@ -93,60 +132,226 @@ end module end defs +variables {R M : Type*} + namespace submodule open module -variables {R M : Type*} - -section torsion_by - -variables [comm_semiring R] [add_comm_monoid M] [module R M] (a : R) +variables [comm_semiring R] [add_comm_monoid M] [module R M] (s : set R) (a : R) @[simp] lemma smul_torsion_by (x : torsion_by R M a) : a • x = 0 := subtype.ext x.prop @[simp] lemma smul_coe_torsion_by (x : torsion_by R M a) : a • (x : M) = 0 := x.prop @[simp] lemma mem_torsion_by_iff (x : M) : x ∈ torsion_by R M a ↔ a • x = 0 := iff.rfl +@[simp] lemma mem_torsion_by_set_iff (x : M) : + x ∈ torsion_by_set R M s ↔ ∀ a : s, (a : R) • x = 0 := +begin + refine ⟨λ h ⟨a, ha⟩, mem_Inf.mp h _ (set.mem_image_of_mem _ ha), λ h, mem_Inf.mpr _⟩, + rintro _ ⟨a, ha, rfl⟩, exact h ⟨a, ha⟩ +end + +@[simp] lemma torsion_by_singleton_eq : torsion_by_set R M {a} = torsion_by R M a := +begin + ext x, + simp only [mem_torsion_by_set_iff, set_coe.forall, subtype.coe_mk, set.mem_singleton_iff, + forall_eq, mem_torsion_by_iff] +end +@[simp] lemma is_torsion_by_singleton_iff : is_torsion_by_set R M {a} ↔ is_torsion_by R M a := +begin + refine ⟨λ h x, @h _ ⟨_, set.mem_singleton _⟩, λ h x, _⟩, + rintro ⟨b, rfl : b = a⟩, exact @h _ +end + +lemma is_torsion_by_set_iff_torsion_by_set_eq_top : + is_torsion_by_set R M s ↔ torsion_by_set R M s = ⊤ := +⟨λ h, eq_top_iff.mpr (λ _ _, (mem_torsion_by_set_iff _ _).mpr $ @h _), + λ h x, by { rw [← mem_torsion_by_set_iff, h], trivial }⟩ + /-- A `a`-torsion module is a module whose `a`-torsion submodule is the full space. -/ lemma is_torsion_by_iff_torsion_by_eq_top : is_torsion_by R M a ↔ torsion_by R M a = ⊤ := -⟨λ h, eq_top_iff.mpr (λ _ _, @h _), λ h x, by { rw [← mem_torsion_by_iff, h], trivial }⟩ +by rw [← torsion_by_singleton_eq, ← is_torsion_by_singleton_iff, + is_torsion_by_set_iff_torsion_by_set_eq_top] + +lemma torsion_by_set_is_torsion_by_set : is_torsion_by_set R (torsion_by_set R M s) s := +λ ⟨x, hx⟩ a, subtype.ext $ (mem_torsion_by_set_iff _ _).mp hx a /-- The `a`-torsion submodule is a `a`-torsion module. -/ lemma torsion_by_is_torsion_by : is_torsion_by R (torsion_by R M a) a := λ _, smul_torsion_by _ _ @[simp] lemma torsion_by_torsion_by_eq_top : torsion_by R (torsion_by R M a) a = ⊤ := (is_torsion_by_iff_torsion_by_eq_top a).mp $ torsion_by_is_torsion_by a +@[simp] lemma torsion_by_set_torsion_by_set_eq_top : + torsion_by_set R (torsion_by_set R M s) s = ⊤ := +(is_torsion_by_set_iff_torsion_by_set_eq_top s).mp $ torsion_by_set_is_torsion_by_set s + +lemma torsion_by_set_le_torsion_by_set_of_subset {s t : set R} (st : s ⊆ t) : + torsion_by_set R M t ≤ torsion_by_set R M s := +Inf_le_Inf $ λ _ ⟨a, ha, h⟩, ⟨a, st ha, h⟩ + +/-- Torsion by a set is torsion by the ideal generated by it. -/ +lemma torsion_by_set_eq_torsion_by_span : + torsion_by_set R M s = torsion_by_set R M (ideal.span s) := +begin + refine le_antisymm (λ x hx, _) (torsion_by_set_le_torsion_by_set_of_subset subset_span), + rw mem_torsion_by_set_iff at hx ⊢, + suffices : ideal.span s ≤ torsion_of R M x, + { rintro ⟨a, ha⟩, exact this ha }, + rw ideal.span_le, exact λ a ha, hx ⟨a, ha⟩ +end +lemma is_torsion_by_set_iff_is_torsion_by_span : + is_torsion_by_set R M s ↔ is_torsion_by_set R M (ideal.span s) := +by rw [is_torsion_by_set_iff_torsion_by_set_eq_top, is_torsion_by_set_iff_torsion_by_set_eq_top, + torsion_by_set_eq_torsion_by_span] + +lemma torsion_by_span_singleton_eq : torsion_by_set R M (R ∙ a) = torsion_by R M a := +((torsion_by_set_eq_torsion_by_span _).symm.trans $ torsion_by_singleton_eq _) +lemma is_torsion_by_span_singleton_iff : is_torsion_by_set R M (R ∙ a) ↔ is_torsion_by R M a := +((is_torsion_by_set_iff_is_torsion_by_span _).symm.trans $ is_torsion_by_singleton_iff _) + +lemma torsion_by_le_torsion_by_of_dvd (a b : R) (dvd : a ∣ b) : + torsion_by R M a ≤ torsion_by R M b := +begin + rw [← torsion_by_span_singleton_eq, ← torsion_by_singleton_eq], + apply torsion_by_set_le_torsion_by_set_of_subset, + rintro c (rfl : c = b), exact ideal.mem_span_singleton.mpr dvd +end + +@[simp] lemma torsion_by_one : torsion_by R M 1 = ⊥ := +eq_bot_iff.mpr (λ _ h, by { rw [mem_torsion_by_iff, one_smul] at h, exact h }) +@[simp] lemma torsion_by_univ : torsion_by_set R M set.univ = ⊥ := +by { rw [eq_bot_iff, ← torsion_by_one, ← torsion_by_singleton_eq], + exact torsion_by_set_le_torsion_by_set_of_subset (λ _ _, trivial) } + +section coprime +open_locale big_operators +open dfinsupp +variables {ι : Type*} {p : ι → R} {S : finset ι} (hp : pairwise (is_coprime on λ s : S, p s)) +include hp + +lemma supr_torsion_by_eq_torsion_by_prod : + (⨆ i : S, torsion_by R M (p i)) = torsion_by R M (∏ i in S, p i) := +begin + cases S.eq_empty_or_nonempty with h h, + { rw [h, finset.prod_empty, torsion_by_one], + convert supr_of_empty _, exact subtype.is_empty_false }, + apply le_antisymm, + { apply supr_le _, rintro ⟨i, is⟩, + exact torsion_by_le_torsion_by_of_dvd _ _ (finset.dvd_prod_of_mem p is) }, + { intros x hx, classical, rw mem_supr_iff_exists_dfinsupp', + cases (exists_sum_eq_one_iff_pairwise_coprime h).mpr hp with f hf, + use equiv_fun_on_fintype.inv_fun (λ i, ⟨(f i * ∏ j in S \ {i}, p j) • x, begin + obtain ⟨i, is⟩ := i, + change p i • (f i * ∏ j in S \ {i}, _) • _ = _, change _ • _ = _ at hx, + rw [smul_smul, mul_comm, mul_assoc, mul_smul, ← finset.prod_eq_prod_diff_singleton_mul is, + hx, smul_zero] + end⟩), + simp only [equiv.inv_fun_as_coe, sum_eq_sum_fintype, coe_eq_zero, eq_self_iff_true, + implies_true_iff, finset.univ_eq_attach, equiv_fun_on_fintype_apply], + change ∑ i : S, ((f i * ∏ j in S \ {i}, p j) • x) = x, + have : ∑ i : S, _ = _ := S.sum_finset_coe (λ i, f i * ∏ j in S \ {i}, p j), + rw [← finset.sum_smul, this, hf, one_smul] } +end -end torsion_by +lemma torsion_by_independent : complete_lattice.independent (λ i : S, torsion_by R M (p i)) := +λ i, begin + classical, + dsimp, rw [disjoint_iff, eq_bot_iff], intros x hx, + rw submodule.mem_inf at hx, obtain ⟨hxi, hxj⟩ := hx, + have hxi : p i • x = 0 := hxi, + rw mem_supr_iff_exists_dfinsupp' at hxj, cases hxj with f hf, + obtain ⟨b, c, h1⟩ := pairwise_coprime_iff_coprime_prod.mp hp i i.2, + rw [mem_bot, ← one_smul _ x, ← h1, add_smul], + convert (zero_add (0:M)), + { rw [mul_smul, hxi, smul_zero] }, + { rw [← hf, smul_sum, sum_eq_zero], + intro j, by_cases ji : j = i, + { convert smul_zero _, + rw ← mem_bot _, convert coe_mem (f j), + symmetry, rw supr_eq_bot, intro hj', + exfalso, exact hj' ji }, + { have hj' : ↑j ∈ S \ {i}, + { rw finset.mem_sdiff, refine ⟨j.2, λ hj', ji _⟩, ext, rw ← finset.mem_singleton, exact hj' }, + rw [finset.prod_eq_prod_diff_singleton_mul hj', ← mul_assoc, mul_smul], + have : (⨆ (H : ¬j = i), torsion_by R M (p j)) ≤ torsion_by R M (p j) := supr_const_le, + have : _ • _ = _ := this (coe_mem _), + rw [this, smul_zero] } } +end +end coprime +end submodule -section quotient +section needs_group +variables [comm_ring R] [add_comm_group M] [module R M] + +namespace submodule +open_locale big_operators +variables {ι : Type*} {p : ι → R} {S : finset ι} (hp : pairwise (is_coprime on λ s : S, p s)) +include hp -variables [comm_ring R] [add_comm_group M] [module R M] (a : R) +/--If the `p i` are pairwise coprime, a `∏ i, p i`-torsion module is the internal direct sum of +its `p i`-torsion submodules.-/ +lemma torsion_is_internal [decidable_eq ι] (hM : torsion_by R M (∏ i in S, p i) = ⊤) : + direct_sum.submodule_is_internal (λ i : S, torsion_by R M (p i)) := +direct_sum.submodule_is_internal_of_independent_of_supr_eq_top + (torsion_by_independent hp) (by { rw ← hM, exact supr_torsion_by_eq_torsion_by_prod hp}) + +end submodule + +namespace module +variables {I : ideal R} (hM : is_torsion_by_set R M I) +include hM -instance : has_scalar (R ⧸ R ∙ a) (torsion_by R M a) := +/-- can't be an instance because hM can't be inferred -/ +def is_torsion_by_set.has_scalar : has_scalar (R ⧸ I) M := { smul := λ b x, quotient.lift_on' b (• x) $ λ b₁ b₂ h, begin - rw submodule.quotient_rel_r_def at h, show b₁ • x = b₂ • x, - obtain ⟨c, h⟩ := ideal.mem_span_singleton'.mp h, - rw [← sub_eq_zero, ← sub_smul, ← h, mul_smul, smul_torsion_by, smul_zero], + have : (-b₁ + b₂) • x = 0 := @hM x ⟨_, h⟩, + rw [add_smul, neg_smul, neg_add_eq_zero] at this, + exact this end } -@[simp] lemma torsion_by.mk_smul (b : R) (x : torsion_by R M a) : - ideal.quotient.mk (R ∙ a) b • x = b • x := rfl +@[simp] lemma is_torsion_by_set.mk_smul (b : R) (x : M) : + by haveI := hM.has_scalar; exact ideal.quotient.mk I b • x = b • x := rfl + +/-- A `(R ⧸ I)`-module is a `R`-module which `is_torsion_by_set R M I`. -/ +def is_torsion_by_set.module : module (R ⧸ I) M := +@function.surjective.module_left _ _ _ _ _ _ _ hM.has_scalar + _ ideal.quotient.mk_surjective (is_torsion_by_set.mk_smul hM) + +end module + +namespace submodule + +instance (I : ideal R) : module (R ⧸ I) (torsion_by_set R M I) := +module.is_torsion_by_set.module $ torsion_by_set_is_torsion_by_set I + +@[simp] lemma torsion_by_set.mk_smul (I : ideal R) (b : R) (x : torsion_by_set R M I) : + ideal.quotient.mk I b • x = b • x := rfl + +instance (I : ideal R) {S : Type*} [has_scalar S R] [has_scalar S M] + [is_scalar_tower S R M] [is_scalar_tower S R R] : + is_scalar_tower S (R ⧸ I) (torsion_by_set R M I) := +{ smul_assoc := λ b d x, quotient.induction_on' d $ λ c, (smul_assoc b c x : _) } /-- The `a`-torsion submodule as a `(R ⧸ R∙a)`-module. -/ -instance : module (R ⧸ R ∙ a) (torsion_by R M a) := -@function.surjective.module_left _ _ (torsion_by R M a) _ _ _ _ _ (ideal.quotient.mk $ R ∙ a) - ideal.quotient.mk_surjective (torsion_by.mk_smul _) +instance (a : R) : module (R ⧸ R ∙ a) (torsion_by R M a) := +module.is_torsion_by_set.module $ + (is_torsion_by_span_singleton_iff a).mpr $ torsion_by_is_torsion_by a + +@[simp] lemma torsion_by.mk_smul (a b : R) (x : torsion_by R M a) : + ideal.quotient.mk (R ∙ a) b • x = b • x := rfl -instance {S : Type*} [has_scalar S R] [has_scalar S M] +instance (a : R) {S : Type*} [has_scalar S R] [has_scalar S M] [is_scalar_tower S R M] [is_scalar_tower S R R] : is_scalar_tower S (R ⧸ R ∙ a) (torsion_by R M a) := { smul_assoc := λ b d x, quotient.induction_on' d $ λ c, (smul_assoc b c x : _) } -end quotient +end submodule +end needs_group +namespace submodule section torsion' +open module variables [comm_semiring R] [add_comm_monoid M] [module R M] variables (S : Type*) [comm_monoid S] [distrib_mul_action S M] [smul_comm_class S R M] @@ -177,6 +382,11 @@ torsion module. -/ /-- The torsion submodule is always a torsion module. -/ lemma torsion_is_torsion : module.is_torsion R (torsion R M) := torsion'_is_torsion' R⁰ +lemma is_torsion'_powers_iff (p : R) : + is_torsion' M (submonoid.powers p) ↔ ∀ x : M, ∃ n : ℕ, p ^ n • x = 0 := +⟨λ h x, let ⟨⟨a, ⟨n, rfl⟩⟩, hx⟩ := @h x in ⟨n, hx⟩, +λ h x, let ⟨n, hn⟩ := h x in ⟨⟨_, ⟨n, rfl⟩⟩, hn⟩⟩ + end torsion' section torsion @@ -228,12 +438,28 @@ instance no_zero_smul_divisors [is_domain R] : no_zero_smul_divisors R (M ⧸ to no_zero_smul_divisors_iff_torsion_eq_bot.mpr torsion_eq_bot end quotient_torsion - end submodule -namespace add_monoid +namespace ideal.quotient + +open submodule -variables {M : Type*} +lemma torsion_by_eq_span_singleton {R : Type*} [comm_ring R] (a b : R) (ha : a ∈ R⁰) : + torsion_by R (R ⧸ R ∙ a * b) a = R ∙ (mk _ b) := +begin + ext x, rw [mem_torsion_by_iff, mem_span_singleton], + obtain ⟨x, rfl⟩ := mk_surjective x, split; intro h, + { rw [← mk_eq_mk, ← quotient.mk_smul, quotient.mk_eq_zero, mem_span_singleton] at h, + obtain ⟨c, h⟩ := h, rw [smul_eq_mul, smul_eq_mul, mul_comm, mul_assoc, + mul_cancel_left_mem_non_zero_divisor ha, mul_comm] at h, + use c, rw [← h, ← mk_eq_mk, ← quotient.mk_smul, smul_eq_mul, mk_eq_mk] }, + { obtain ⟨c, h⟩ := h, + rw [← h, smul_comm, ← mk_eq_mk, ← quotient.mk_smul, + (quotient.mk_eq_zero _).mpr $ mem_span_singleton_self _, smul_zero] } +end +end ideal.quotient + +namespace add_monoid theorem is_torsion_iff_is_torsion_nat [add_comm_monoid M] : add_monoid.is_torsion M ↔ module.is_torsion ℕ M := diff --git a/src/ring_theory/coprime/lemmas.lean b/src/ring_theory/coprime/lemmas.lean index 12da89c664589..454d672e59b7b 100644 --- a/src/ring_theory/coprime/lemmas.lean +++ b/src/ring_theory/coprime/lemmas.lean @@ -22,7 +22,9 @@ universes u v variables {R : Type u} {I : Type v} [comm_semiring R] {x y z : R} {s : I → R} {t : finset I} -open_locale big_operators classical +open_locale big_operators +section +open_locale classical theorem nat.is_coprime_iff_coprime {m n : ℕ} : is_coprime (m : ℤ) n ↔ nat.coprime m n := ⟨λ ⟨a, b, H⟩, nat.eq_one_of_dvd_one $ int.coe_nat_dvd.1 $ by { rw [int.coe_nat_one, ← H], @@ -71,9 +73,11 @@ theorem fintype.prod_dvd_of_coprime [fintype I] (Hs : pairwise (is_coprime on s) (Hs1 : ∀ i, s i ∣ z) : ∏ x, s x ∣ z := finset.prod_dvd_of_coprime (Hs.set_pairwise _) (λ i _, Hs1 i) +end + open finset -lemma exists_sum_eq_one_iff_pairwise_coprime (h : t.nonempty) : +lemma exists_sum_eq_one_iff_pairwise_coprime [decidable_eq I] (h : t.nonempty) : (∃ μ : I → R, ∑ i in t, μ i * ∏ j in t \ {i}, s j = 1) ↔ pairwise (is_coprime on λ i : t, s i) := begin refine h.cons_induction _ _; clear' t h, @@ -118,7 +122,7 @@ begin convert sdiff_sdiff_comm, rw [sdiff_singleton_eq_erase, erase_insert hat] } end -lemma exists_sum_eq_one_iff_pairwise_coprime' [fintype I] [nonempty I] : +lemma exists_sum_eq_one_iff_pairwise_coprime' [fintype I] [nonempty I] [decidable_eq I] : (∃ μ : I → R, ∑ (i : I), μ i * ∏ j in {i}ᶜ, s j = 1) ↔ pairwise (is_coprime on s) := begin convert exists_sum_eq_one_iff_pairwise_coprime finset.univ_nonempty using 1, @@ -126,7 +130,7 @@ begin assumption end -lemma pairwise_coprime_iff_coprime_prod : +lemma pairwise_coprime_iff_coprime_prod [decidable_eq I] : pairwise (is_coprime on λ i : t, s i) ↔ ∀ i ∈ t, is_coprime (s i) (∏ j in t \ {i}, (s j)) := begin refine ⟨λ hp i hi, is_coprime.prod_right_iff.mpr (λ j hj, _), λ hp, _⟩, From e1f00bc7f9533db84552999828c627155ee96075 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 4 May 2022 11:10:43 +0000 Subject: [PATCH 094/153] feat(order/well_founded): Well founded relations are asymmetric and irreflexive (#13692) Co-authored-by: Eric Wieser --- src/order/well_founded.lean | 19 +++++++++++++++++++ 1 file changed, 19 insertions(+) diff --git a/src/order/well_founded.lean b/src/order/well_founded.lean index 347857614d4ad..2642c4ba58431 100644 --- a/src/order/well_founded.lean +++ b/src/order/well_founded.lean @@ -20,6 +20,25 @@ and provide a few new definitions: `well_founded.min`, `well_founded.sup`, and ` variables {α : Type*} namespace well_founded + +theorem not_gt_of_lt {α : Sort*} {r : α → α → Prop} (h : well_founded r) : + ∀ ⦃a b⦄, r a b → ¬ r b a +| a := λ b hab hba, not_gt_of_lt hba hab +using_well_founded { rel_tac := λ _ _, `[exact ⟨_, h⟩], + dec_tac := tactic.assumption } + +protected theorem is_asymm {α : Sort*} {r : α → α → Prop} (h : well_founded r) : is_asymm α r := +⟨h.not_gt_of_lt⟩ + +instance {α : Sort*} [has_well_founded α] : is_asymm α has_well_founded.r := +has_well_founded.wf.is_asymm + +protected theorem is_irrefl {α : Sort*} {r : α → α → Prop} (h : well_founded r) : is_irrefl α r := +(@is_asymm.is_irrefl α r h.is_asymm) + +instance {α : Sort*} [has_well_founded α] : is_irrefl α has_well_founded.r := +is_asymm.is_irrefl + /-- If `r` is a well-founded relation, then any nonempty set has a minimal element with respect to `r`. -/ theorem has_min {α} {r : α → α → Prop} (H : well_founded r) From 402e5647e15c7ffc54642586a4643653809fea42 Mon Sep 17 00:00:00 2001 From: antoinelab01 Date: Wed, 4 May 2022 11:10:44 +0000 Subject: [PATCH 095/153] feat(linear_algebra/prod): linear version of prod_map (#13751) Co-authored-by: antoinelab01 <66086247+antoinelab01@users.noreply.github.com> --- src/linear_algebra/prod.lean | 47 +++++++++++++++++++++++++++++++++--- 1 file changed, 44 insertions(+), 3 deletions(-) diff --git a/src/linear_algebra/prod.lean b/src/linear_algebra/prod.lean index f1dead99d12d7..b2ca23c5c34c1 100644 --- a/src/linear_algebra/prod.lean +++ b/src/linear_algebra/prod.lean @@ -5,6 +5,7 @@ Authors: Johannes Hölzl, Mario Carneiro, Kevin Buzzard, Yury Kudryashov, Eric W -/ import linear_algebra.span import order.partial_sups +import algebra.algebra.basic /-! ### Products of modules @@ -250,12 +251,37 @@ lemma prod_map_comp (f₁₂ : M →ₗ[R] M₂) (f₂₃ : M₂ →ₗ[R] M₃) lemma prod_map_mul (f₁₂ : M →ₗ[R] M) (f₂₃ : M →ₗ[R] M) (g₁₂ : M₂ →ₗ[R] M₂) (g₂₃ : M₂ →ₗ[R] M₂) : f₂₃.prod_map g₂₃ * f₁₂.prod_map g₁₂ = (f₂₃ * f₁₂).prod_map (g₂₃ * g₁₂) := rfl -/-- `linear_map.prod_map` as a `monoid_hom` -/ +lemma prod_map_add (f₁ : M →ₗ[R] M₃) (f₂ : M →ₗ[R] M₃) (g₁ : M₂ →ₗ[R] M₄) (g₂ : M₂ →ₗ[R] M₄) : + (f₁ + f₂).prod_map (g₁ + g₂) = f₁.prod_map g₁ + f₂.prod_map g₂ := rfl + +@[simp] lemma prod_map_zero : + (0 : M →ₗ[R] M₂).prod_map (0 : M₃ →ₗ[R] M₄) = 0 := rfl + +@[simp] lemma prod_map_smul + [module S M₃] [module S M₄] [smul_comm_class R S M₃] [smul_comm_class R S M₄] + (s : S) (f : M →ₗ[R] M₃) (g : M₂ →ₗ[R] M₄) : prod_map (s • f) (s • g) = s • prod_map f g := rfl + +variables (R M M₂ M₃ M₄) + +/-- `linear_map.prod_map` as a `linear_map` -/ +@[simps] +def prod_map_linear + [module S M₃] [module S M₄] [smul_comm_class R S M₃] [smul_comm_class R S M₄] : + ((M →ₗ[R] M₃) × (M₂ →ₗ[R] M₄)) →ₗ[S] ((M × M₂) →ₗ[R] (M₃ × M₄)) := +{ to_fun := λ f, prod_map f.1 f.2, + map_add' := λ _ _, rfl, + map_smul' := λ _ _, rfl} + +/-- `linear_map.prod_map` as a `ring_hom` -/ @[simps] -def prod_map_monoid_hom : (M →ₗ[R] M) × (M₂ →ₗ[R] M₂) →* ((M × M₂) →ₗ[R] (M × M₂)) := +def prod_map_ring_hom : (M →ₗ[R] M) × (M₂ →ₗ[R] M₂) →+* ((M × M₂) →ₗ[R] (M × M₂)) := { to_fun := λ f, prod_map f.1 f.2, map_one' := prod_map_one, - map_mul' := λ _ _, prod_map_mul _ _ _ _ } + map_zero' := rfl, + map_add' := λ _ _, rfl, + map_mul' := λ _ _, rfl } + +variables {R M M₂ M₃ M₄} section map_mul @@ -276,6 +302,21 @@ end linear_map end prod +namespace linear_map + +variables (R M M₂) + +variables [comm_semiring R] +variables [add_comm_monoid M] [add_comm_monoid M₂] +variables [module R M] [module R M₂] + +/-- `linear_map.prod_map` as an `algebra_hom` -/ +@[simps] +def prod_map_alg_hom : (module.End R M) × (module.End R M₂) →ₐ[R] module.End R (M × M₂) := +{ commutes' := λ _, rfl, ..prod_map_ring_hom R M M₂ } + +end linear_map + namespace linear_map open submodule From a057441eb14fe0e65a5588084e7e2c1193b82579 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 4 May 2022 11:10:45 +0000 Subject: [PATCH 096/153] feat(order/basic): Notation for `order_dual` (#13798) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define `αᵒᵈ` as notation for `order_dual α` and replace current uses. [Zulip poll](https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/Notation.20for.20order_dual/near/280629129) --- src/algebra/big_operators/multiset.lean | 2 +- src/algebra/big_operators/order.lean | 6 +- src/algebra/bounds.lean | 4 +- src/algebra/group_power/order.lean | 11 +- src/algebra/indicator_function.lean | 4 +- src/algebra/order/archimedean.lean | 3 +- src/algebra/order/field.lean | 4 +- src/algebra/order/group.lean | 28 ++-- src/algebra/order/module.lean | 6 +- src/algebra/order/monoid.lean | 80 +++++----- src/algebra/order/ring.lean | 28 ++-- src/algebra/order/smul.lean | 20 +-- src/algebra/order/with_zero.lean | 8 +- src/algebra/support.lean | 2 +- .../prime_spectrum/basic.lean | 6 +- .../projective_spectrum/topology.lean | 6 +- .../box_integral/partition/filter.lean | 4 +- src/analysis/convex/basic.lean | 18 +-- src/analysis/convex/extrema.lean | 5 +- src/analysis/convex/function.lean | 18 +-- src/analysis/convex/jensen.lean | 8 +- src/analysis/convex/quasiconvex.lean | 4 +- src/analysis/convex/strict.lean | 3 +- src/analysis/inner_product_space/basic.lean | 2 +- .../normed_space/lattice_ordered_group.lean | 6 +- src/combinatorics/pigeonhole.lean | 12 +- src/data/fin/basic.lean | 2 +- src/data/fin/tuple/basic.lean | 2 +- src/data/finset/lattice.lean | 147 ++++++++---------- src/data/finset/locally_finite.lean | 4 +- src/data/finset/pi_induction.lean | 2 +- src/data/finset/sigma.lean | 2 +- src/data/fintype/basic.lean | 12 +- src/data/list/big_operators.lean | 2 +- src/data/list/min_max.lean | 35 ++--- src/data/nat/lattice.lean | 4 +- src/data/nat/pairing.lean | 2 +- src/data/ordmap/ordset.lean | 20 +-- src/data/real/ennreal.lean | 8 +- src/data/real/ereal.lean | 2 +- src/data/set/finite.lean | 7 +- src/data/set/function.lean | 4 +- src/data/set/intervals/basic.lean | 7 +- src/data/set/intervals/infinite.lean | 3 +- src/data/set/intervals/pi.lean | 2 +- src/data/set/intervals/with_bot_top.lean | 4 +- src/data/sum/order.lean | 6 +- src/field_theory/galois.lean | 2 +- src/linear_algebra/affine_space/ordered.lean | 20 +-- src/linear_algebra/basic.lean | 2 +- src/linear_algebra/prod.lean | 2 +- .../constructions/borel_space.lean | 16 +- src/measure_theory/function/ess_sup.lean | 12 +- .../function/locally_integrable.lean | 6 +- .../integral/interval_integral.lean | 2 +- src/measure_theory/lattice.lean | 12 +- src/measure_theory/measurable_space_def.lean | 2 +- src/measure_theory/measure/measure_space.lean | 3 +- src/measure_theory/pi_system.lean | 2 +- src/order/antisymmetrization.lean | 2 +- src/order/atoms.lean | 26 ++-- src/order/basic.lean | 39 ++--- src/order/boolean_algebra.lean | 2 +- src/order/bounded.lean | 30 ++-- src/order/bounded_order.lean | 22 +-- src/order/bounds.lean | 59 ++++--- src/order/category/BoolAlg.lean | 2 +- src/order/category/BoundedDistribLattice.lean | 2 +- src/order/category/BoundedLattice.lean | 2 +- src/order/category/BoundedOrder.lean | 2 +- src/order/category/CompleteLattice.lean | 2 +- src/order/category/DistribLattice.lean | 2 +- src/order/category/FinBoolAlg.lean | 2 +- src/order/category/FinPartialOrder.lean | 2 +- src/order/category/Lattice.lean | 3 +- src/order/category/LinearOrder.lean | 2 +- src/order/category/NonemptyFinLinOrd.lean | 4 +- src/order/category/PartialOrder.lean | 2 +- src/order/category/Preorder.lean | 2 +- src/order/category/Semilattice.lean | 4 +- src/order/circular.lean | 22 +-- src/order/compare.lean | 4 +- src/order/complete_boolean_algebra.lean | 30 ++-- src/order/complete_lattice.lean | 119 +++++++------- src/order/concept.lean | 2 +- src/order/conditionally_complete_lattice.lean | 77 +++++---- src/order/cover.lean | 5 +- src/order/directed.lean | 6 +- src/order/filter/at_top_bot.lean | 129 ++++++++------- src/order/filter/basic.lean | 2 +- src/order/filter/cofinite.lean | 4 +- src/order/filter/extr.lean | 7 +- src/order/fixed_points.lean | 6 +- src/order/grade.lean | 16 +- src/order/hom/basic.lean | 36 ++--- src/order/hom/bounded.lean | 16 +- src/order/hom/complete_lattice.lean | 16 +- src/order/hom/lattice.lean | 31 ++-- src/order/iterate.lean | 8 +- src/order/lattice.lean | 61 ++++---- src/order/liminf_limsup.lean | 40 ++--- src/order/locally_finite.lean | 6 +- src/order/max.lean | 22 ++- src/order/min_max.lean | 19 +-- src/order/modular_lattice.lean | 10 +- src/order/monotone.lean | 13 +- src/order/order_dual.lean | 59 +++---- src/order/pfilter.lean | 6 +- src/order/rel_classes.lean | 6 +- src/order/rel_iso.lean | 2 +- src/order/succ_pred/basic.lean | 22 ++- src/order/succ_pred/relation.lean | 14 +- src/order/upper_lower.lean | 8 +- src/order/well_founded.lean | 2 +- src/order/well_founded_set.lean | 3 +- src/order/zorn.lean | 5 +- src/ring_theory/artinian.lean | 8 +- src/ring_theory/nullstellensatz.lean | 2 +- src/ring_theory/valuation/basic.lean | 9 +- .../valuation/valuation_subring.lean | 3 +- src/topology/algebra/order/basic.lean | 111 ++++++------- src/topology/algebra/order/compact.lean | 27 ++-- src/topology/algebra/order/extend_from.lean | 2 +- .../algebra/order/intermediate_value.lean | 6 +- src/topology/algebra/order/left_right.lean | 2 +- src/topology/algebra/order/liminf_limsup.lean | 10 +- .../algebra/order/monotone_continuity.lean | 7 +- .../algebra/order/monotone_convergence.lean | 49 +++--- src/topology/continuous_function/ordered.lean | 4 +- src/topology/local_extr.lean | 2 +- src/topology/order.lean | 24 ++- src/topology/order/lattice.lean | 8 +- src/topology/semicontinuous.lean | 28 ++-- 133 files changed, 909 insertions(+), 1096 deletions(-) diff --git a/src/algebra/big_operators/multiset.lean b/src/algebra/big_operators/multiset.lean index 19a9cf026dff9..fdeecd970c4e6 100644 --- a/src/algebra/big_operators/multiset.lean +++ b/src/algebra/big_operators/multiset.lean @@ -318,7 +318,7 @@ prod_le_prod_of_rel_le $ rel_map_left.2 $ rel_refl_of_refl_on h @[to_additive] lemma prod_le_sum_prod (f : α → α) (h : ∀ x, x ∈ s → x ≤ f x) : s.prod ≤ (s.map f).prod := -@prod_map_le_prod (order_dual α) _ _ f h +@prod_map_le_prod αᵒᵈ _ _ f h @[to_additive card_nsmul_le_sum] lemma pow_card_le_prod (h : ∀ x ∈ s, a ≤ x) : a ^ s.card ≤ s.prod := diff --git a/src/algebra/big_operators/order.lean b/src/algebra/big_operators/order.lean index 08bad25be2983..8c7030a2ecf03 100644 --- a/src/algebra/big_operators/order.lean +++ b/src/algebra/big_operators/order.lean @@ -161,7 +161,7 @@ end @[to_additive sum_eq_zero_iff_of_nonneg] lemma prod_eq_one_iff_of_le_one' : (∀ i ∈ s, f i ≤ 1) → (∏ i in s, f i = 1 ↔ ∀ i ∈ s, f i = 1) := -@prod_eq_one_iff_of_one_le' _ (order_dual N) _ _ _ +@prod_eq_one_iff_of_one_le' _ Nᵒᵈ _ _ _ @[to_additive single_le_sum] lemma single_le_prod' (hf : ∀ i ∈ s, 1 ≤ f i) {a} (h : a ∈ s) : f a ≤ (∏ x in s, f x) := @@ -181,7 +181,7 @@ end @[to_additive card_nsmul_le_sum] lemma pow_card_le_prod (s : finset ι) (f : ι → N) (n : N) (h : ∀ x ∈ s, n ≤ f x) : n ^ s.card ≤ s.prod f := -@finset.prod_le_pow_card _ (order_dual N) _ _ _ _ h +@finset.prod_le_pow_card _ Nᵒᵈ _ _ _ _ h lemma card_bUnion_le_card_mul [decidable_eq β] (s : finset ι) (f : ι → finset β) (n : ℕ) (h : ∀ a ∈ s, (f a).card ≤ n) : @@ -204,7 +204,7 @@ calc (∏ y in t, ∏ x in s.filter (λ x, g x = y), f x) ≤ lemma prod_le_prod_fiberwise_of_prod_fiber_le_one' {t : finset ι'} {g : ι → ι'} {f : ι → N} (h : ∀ y ∉ t, (∏ x in s.filter (λ x, g x = y), f x) ≤ 1) : (∏ x in s, f x) ≤ ∏ y in t, ∏ x in s.filter (λ x, g x = y), f x := -@prod_fiberwise_le_prod_of_one_le_prod_fiber' _ (order_dual N) _ _ _ _ _ _ _ h +@prod_fiberwise_le_prod_of_one_le_prod_fiber' _ Nᵒᵈ _ _ _ _ _ _ _ h end ordered_comm_monoid diff --git a/src/algebra/bounds.lean b/src/algebra/bounds.lean index 4ee7d179402ec..50eab7e89589c 100644 --- a/src/algebra/bounds.lean +++ b/src/algebra/bounds.lean @@ -69,11 +69,11 @@ image2_subset_iff.2 $ λ x hx y hy, mul_mem_upper_bounds_mul hx hy @[to_additive] lemma mul_mem_lower_bounds_mul {s t : set M} {a b : M} (ha : a ∈ lower_bounds s) (hb : b ∈ lower_bounds t) : a * b ∈ lower_bounds (s * t) := -@mul_mem_upper_bounds_mul (order_dual M) _ _ _ _ _ _ _ _ ha hb +@mul_mem_upper_bounds_mul Mᵒᵈ _ _ _ _ _ _ _ _ ha hb @[to_additive] lemma subset_lower_bounds_mul (s t : set M) : lower_bounds s * lower_bounds t ⊆ lower_bounds (s * t) := -@subset_upper_bounds_mul (order_dual M) _ _ _ _ _ _ +@subset_upper_bounds_mul Mᵒᵈ _ _ _ _ _ _ @[to_additive] lemma bdd_above.mul {s t : set M} (hs : bdd_above s) (ht : bdd_above t) : bdd_above (s * t) := diff --git a/src/algebra/group_power/order.lean b/src/algebra/group_power/order.lean index 04d0ff5c15b5f..aa2ab3534f216 100644 --- a/src/algebra/group_power/order.lean +++ b/src/algebra/group_power/order.lean @@ -34,8 +34,7 @@ theorem one_le_pow_of_one_le' {a : M} (H : 1 ≤ a) : ∀ n : ℕ, 1 ≤ a ^ n | (k + 1) := by { rw pow_succ, exact one_le_mul H (one_le_pow_of_one_le' k) } @[to_additive nsmul_nonpos] -theorem pow_le_one' {a : M} (H : a ≤ 1) (n : ℕ) : a ^ n ≤ 1 := -@one_le_pow_of_one_le' (order_dual M) _ _ _ _ H n +lemma pow_le_one' {a : M} (H : a ≤ 1) (n : ℕ) : a ^ n ≤ 1 := @one_le_pow_of_one_le' Mᵒᵈ _ _ _ _ H n @[to_additive nsmul_le_nsmul] theorem pow_le_pow' {a : M} {n m : ℕ} (ha : 1 ≤ a) (h : n ≤ m) : a ^ n ≤ a ^ m := @@ -45,7 +44,7 @@ calc a ^ n ≤ a ^ n * a ^ k : le_mul_of_one_le_right' (one_le_pow_of_one_le' ha @[to_additive nsmul_le_nsmul_of_nonpos] theorem pow_le_pow_of_le_one' {a : M} {n m : ℕ} (ha : a ≤ 1) (h : n ≤ m) : a ^ m ≤ a ^ n := -@pow_le_pow' (order_dual M) _ _ _ _ _ _ ha h +@pow_le_pow' Mᵒᵈ _ _ _ _ _ _ ha h @[to_additive nsmul_pos] theorem one_lt_pow' {a : M} (ha : 1 < a) {k : ℕ} (hk : k ≠ 0) : 1 < a ^ k := @@ -59,8 +58,8 @@ begin end @[to_additive nsmul_neg] -theorem pow_lt_one' {a : M} (ha : a < 1) {k : ℕ} (hk : k ≠ 0) : a ^ k < 1 := -@one_lt_pow' (order_dual M) _ _ _ _ ha k hk +lemma pow_lt_one' {a : M} (ha : a < 1) {k : ℕ} (hk : k ≠ 0) : a ^ k < 1 := +@one_lt_pow' Mᵒᵈ _ _ _ _ ha k hk @[to_additive nsmul_lt_nsmul] theorem pow_lt_pow' [covariant_class M M (*) (<)] {a : M} {n m : ℕ} (ha : 1 < a) (h : n < m) : @@ -88,7 +87,7 @@ lemma one_le_pow_iff {x : M} {n : ℕ} (hn : n ≠ 0) : 1 ≤ x ^ n ↔ 1 ≤ x @[to_additive] lemma pow_le_one_iff {x : M} {n : ℕ} (hn : n ≠ 0) : x ^ n ≤ 1 ↔ x ≤ 1 := -@one_le_pow_iff (order_dual M) _ _ _ _ _ hn +@one_le_pow_iff Mᵒᵈ _ _ _ _ _ hn @[to_additive nsmul_pos_iff] lemma one_lt_pow_iff {x : M} {n : ℕ} (hn : n ≠ 0) : 1 < x ^ n ↔ 1 < x := diff --git a/src/algebra/indicator_function.lean b/src/algebra/indicator_function.lean index 2dffdf02a2538..2762a65b1b36b 100644 --- a/src/algebra/indicator_function.lean +++ b/src/algebra/indicator_function.lean @@ -482,7 +482,7 @@ end @[to_additive] lemma le_mul_indicator_apply {y} (hfg : a ∈ s → y ≤ g a) (hf : a ∉ s → y ≤ 1) : y ≤ mul_indicator s g a := -@mul_indicator_apply_le' α (order_dual M) ‹_› _ _ _ _ _ hfg hf +@mul_indicator_apply_le' α Mᵒᵈ ‹_› _ _ _ _ _ hfg hf @[to_additive] lemma le_mul_indicator (hfg : ∀ a ∈ s, f a ≤ g a) (hf : ∀ a ∉ s, f a ≤ 1) : f ≤ mul_indicator s g := @@ -573,7 +573,7 @@ end lemma indicator_nonpos_le_indicator {β} [linear_order β] [has_zero β] (s : set α) (f : α → β) : {x | f x ≤ 0}.indicator f ≤ s.indicator f := -@indicator_le_indicator_nonneg α (order_dual β) _ _ s f +@indicator_le_indicator_nonneg α βᵒᵈ _ _ s f end set diff --git a/src/algebra/order/archimedean.lean b/src/algebra/order/archimedean.lean index 1dfd6928e874a..3a307b1cb2281 100644 --- a/src/algebra/order/archimedean.lean +++ b/src/algebra/order/archimedean.lean @@ -37,8 +37,7 @@ such that `0 < y` there exists a natural number `n` such that `x ≤ n • y`. - class archimedean (α) [ordered_add_comm_monoid α] : Prop := (arch : ∀ (x : α) {y}, 0 < y → ∃ n : ℕ, x ≤ n • y) -instance order_dual.archimedean [ordered_add_comm_group α] [archimedean α] : - archimedean (order_dual α) := +instance order_dual.archimedean [ordered_add_comm_group α] [archimedean α] : archimedean αᵒᵈ := ⟨λ x y hy, let ⟨n, hn⟩ := archimedean.arch (-x : α) (neg_pos.2 hy) in ⟨n, by rwa [neg_nsmul, neg_le_neg_iff] at hn⟩⟩ diff --git a/src/algebra/order/field.lean b/src/algebra/order/field.lean index 28f70e4a0629e..a2c1521c88c89 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 α (order_dual α) _ _ _ _ _ (λ x y, div_le_div_of_nonpos_of_le hc) +eq.symm $ @monotone.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 α (order_dual α) _ _ _ _ _ (λ x y, div_le_div_of_nonpos_of_le hc) +eq.symm $ @monotone.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 diff --git a/src/algebra/order/group.lean b/src/algebra/order/group.lean index ef151b9840a5d..af87bcf29f2f8 100644 --- a/src/algebra/order/group.lean +++ b/src/algebra/order/group.lean @@ -91,14 +91,14 @@ instance ordered_comm_group.has_exists_mul_of_le (α : Type u) has_exists_mul_of_le α := ⟨λ a b hab, ⟨b * a⁻¹, (mul_inv_cancel_comm_assoc a b).symm⟩⟩ -@[to_additive] instance [h : has_inv α] : has_inv (order_dual α) := h -@[to_additive] instance [h : has_div α] : has_div (order_dual α) := h -@[to_additive] instance [h : has_involutive_inv α] : has_involutive_inv (order_dual α) := h -@[to_additive] instance [h : div_inv_monoid α] : div_inv_monoid (order_dual α) := h -@[to_additive] instance [h : group α] : group (order_dual α) := h -@[to_additive] instance [h : comm_group α] : comm_group (order_dual α) := h - -@[to_additive] instance [ordered_comm_group α] : ordered_comm_group (order_dual α) := +@[to_additive] instance [h : has_inv α] : has_inv αᵒᵈ := h +@[to_additive] instance [h : has_div α] : has_div αᵒᵈ := h +@[to_additive] instance [h : has_involutive_inv α] : has_involutive_inv αᵒᵈ := h +@[to_additive] instance [h : div_inv_monoid α] : div_inv_monoid αᵒᵈ := h +@[to_additive] instance [h : group α] : group αᵒᵈ := h +@[to_additive] instance [h : comm_group α] : comm_group αᵒᵈ := h + +@[to_additive] instance [ordered_comm_group α] : ordered_comm_group αᵒᵈ := { .. order_dual.ordered_comm_monoid, .. order_dual.group } section group @@ -292,7 +292,7 @@ variable (α) /-- `x ↦ x⁻¹` as an order-reversing equivalence. -/ @[to_additive "`x ↦ -x` as an order-reversing equivalence.", simps] -def order_iso.inv : α ≃o order_dual α := +def order_iso.inv : α ≃o αᵒᵈ := { to_equiv := (equiv.inv α).trans order_dual.to_dual, map_rel_iff' := λ a b, @inv_le_inv_iff α _ _ _ _ _ _ } @@ -853,7 +853,7 @@ calc a ≤ b * (b⁻¹ * c) : h _ (lt_inv_mul_iff_lt.mpr hc) @[to_additive] lemma le_of_forall_lt_one_mul_le (h : ∀ ε < 1, a * ε ≤ b) : a ≤ b := -@le_of_forall_one_lt_le_mul (order_dual α) _ _ _ _ _ _ h +@le_of_forall_one_lt_le_mul αᵒᵈ _ _ _ _ _ _ h @[to_additive] lemma le_of_forall_one_lt_div_le (h : ∀ ε : α, 1 < ε → a / ε ≤ b) : a ≤ b := @@ -866,7 +866,7 @@ lemma le_iff_forall_one_lt_le_mul : a ≤ b ↔ ∀ ε, 1 < ε → a ≤ b * ε @[to_additive] lemma le_iff_forall_lt_one_mul_le : a ≤ b ↔ ∀ ε < 1, a * ε ≤ b := -@le_iff_forall_one_lt_le_mul (order_dual α) _ _ _ _ _ _ +@le_iff_forall_one_lt_le_mul αᵒᵈ _ _ _ _ _ _ end densely_ordered @@ -897,7 +897,7 @@ multiplication is monotone. -/ class linear_ordered_comm_group (α : Type u) extends ordered_comm_group α, linear_order α @[to_additive] instance [linear_ordered_comm_group α] : - linear_ordered_comm_group (order_dual α) := + linear_ordered_comm_group αᵒᵈ := { .. order_dual.ordered_comm_group, .. order_dual.linear_order α } section linear_ordered_comm_group @@ -933,11 +933,11 @@ mul_lt_mul_left' h c @[to_additive min_neg_neg] lemma min_inv_inv' (a b : α) : min (a⁻¹) (b⁻¹) = (max a b)⁻¹ := -eq.symm $ @monotone.map_max α (order_dual α) _ _ has_inv.inv a b $ λ a b, inv_le_inv_iff.mpr +eq.symm $ @monotone.map_max α αᵒᵈ _ _ has_inv.inv a b $ λ a b, inv_le_inv_iff.mpr @[to_additive max_neg_neg] lemma max_inv_inv' (a b : α) : max (a⁻¹) (b⁻¹) = (min a b)⁻¹ := -eq.symm $ @monotone.map_min α (order_dual α) _ _ has_inv.inv a b $ λ a b, inv_le_inv_iff.mpr +eq.symm $ @monotone.map_min α αᵒᵈ _ _ has_inv.inv a b $ λ a b, inv_le_inv_iff.mpr @[to_additive min_sub_sub_right] lemma min_div_div_right' (a b c : α) : min (a / c) (b / c) = min a b / c := diff --git a/src/algebra/order/module.lean b/src/algebra/order/module.lean index 45cf0ea3e7b90..6c236a59482b4 100644 --- a/src/algebra/order/module.lean +++ b/src/algebra/order/module.lean @@ -29,7 +29,7 @@ variables {k M N : Type*} namespace order_dual -instance [semiring k] [ordered_add_comm_monoid M] [module k M] : module k (order_dual M) := +instance [semiring k] [ordered_add_comm_monoid M] [module k M] : module k Mᵒᵈ := { add_smul := λ r s x, order_dual.rec (add_smul _ _) x, zero_smul := λ m, order_dual.rec (zero_smul _) m } @@ -110,7 +110,7 @@ calc ... = 0 : smul_zero' M c lemma smul_nonneg_of_nonpos_of_nonpos (hc : c ≤ 0) (ha : a ≤ 0) : 0 ≤ c • a := -@smul_nonpos_of_nonpos_of_nonneg k (order_dual M) _ _ _ _ _ _ hc ha +@smul_nonpos_of_nonpos_of_nonneg k Mᵒᵈ _ _ _ _ _ _ hc ha alias smul_pos_iff_of_neg ↔ _ smul_pos_of_neg_of_neg alias smul_neg_iff_of_pos ↔ _ smul_neg_of_pos_of_neg @@ -149,7 +149,7 @@ end variables (M) /-- Left scalar multiplication as an order isomorphism. -/ -@[simps] def order_iso.smul_left_dual {c : k} (hc : c < 0) : M ≃o order_dual M := +@[simps] def order_iso.smul_left_dual {c : k} (hc : c < 0) : M ≃o Mᵒᵈ := { to_fun := λ b, order_dual.to_dual (c • b), inv_fun := λ b, c⁻¹ • (order_dual.of_dual b), left_inv := inv_smul_smul₀ hc.ne, diff --git a/src/algebra/order/monoid.lean b/src/algebra/order/monoid.lean index 66f35e177ee4b..f03f7d9a1c43b 100644 --- a/src/algebra/order/monoid.lean +++ b/src/algebra/order/monoid.lean @@ -672,90 +672,90 @@ end linear_ordered_cancel_comm_monoid namespace order_dual -@[to_additive] instance [h : has_mul α] : has_mul (order_dual α) := h -@[to_additive] instance [h : has_one α] : has_one (order_dual α) := h -@[to_additive] instance [h : semigroup α] : semigroup (order_dual α) := h -@[to_additive] instance [h : comm_semigroup α] : comm_semigroup (order_dual α) := h -@[to_additive] instance [h : mul_one_class α] : mul_one_class (order_dual α) := h -@[to_additive] instance [h : monoid α] : monoid (order_dual α) := h -@[to_additive] instance [h : comm_monoid α] : comm_monoid (order_dual α) := h -@[to_additive] instance [h : left_cancel_monoid α] : left_cancel_monoid (order_dual α) := h -@[to_additive] instance [h : right_cancel_monoid α] : right_cancel_monoid (order_dual α) := h -@[to_additive] instance [h : cancel_monoid α] : cancel_monoid (order_dual α) := h -@[to_additive] instance [h : cancel_comm_monoid α] : cancel_comm_monoid (order_dual α) := h -instance [h : mul_zero_class α] : mul_zero_class (order_dual α) := h -instance [h : mul_zero_one_class α] : mul_zero_one_class (order_dual α) := h -instance [h : monoid_with_zero α] : monoid_with_zero (order_dual α) := h -instance [h : comm_monoid_with_zero α] : comm_monoid_with_zero (order_dual α) := h -instance [h : cancel_comm_monoid_with_zero α] : cancel_comm_monoid_with_zero (order_dual α) := h +@[to_additive] instance [h : has_mul α] : has_mul αᵒᵈ := h +@[to_additive] instance [h : has_one α] : has_one αᵒᵈ := h +@[to_additive] instance [h : semigroup α] : semigroup αᵒᵈ := h +@[to_additive] instance [h : comm_semigroup α] : comm_semigroup αᵒᵈ := h +@[to_additive] instance [h : mul_one_class α] : mul_one_class αᵒᵈ := h +@[to_additive] instance [h : monoid α] : monoid αᵒᵈ := h +@[to_additive] instance [h : comm_monoid α] : comm_monoid αᵒᵈ := h +@[to_additive] instance [h : left_cancel_monoid α] : left_cancel_monoid αᵒᵈ := h +@[to_additive] instance [h : right_cancel_monoid α] : right_cancel_monoid αᵒᵈ := h +@[to_additive] instance [h : cancel_monoid α] : cancel_monoid αᵒᵈ := h +@[to_additive] instance [h : cancel_comm_monoid α] : cancel_comm_monoid αᵒᵈ := h +instance [h : mul_zero_class α] : mul_zero_class αᵒᵈ := h +instance [h : mul_zero_one_class α] : mul_zero_one_class αᵒᵈ := h +instance [h : monoid_with_zero α] : monoid_with_zero αᵒᵈ := h +instance [h : comm_monoid_with_zero α] : comm_monoid_with_zero αᵒᵈ := h +instance [h : cancel_comm_monoid_with_zero α] : cancel_comm_monoid_with_zero αᵒᵈ := h @[to_additive] instance contravariant_class_mul_le [has_le α] [has_mul α] [c : contravariant_class α α (*) (≤)] : - contravariant_class (order_dual α) (order_dual α) (*) (≤) := + contravariant_class αᵒᵈ αᵒᵈ (*) (≤) := ⟨c.1.flip⟩ @[to_additive] instance covariant_class_mul_le [has_le α] [has_mul α] [c : covariant_class α α (*) (≤)] : - covariant_class (order_dual α) (order_dual α) (*) (≤) := + covariant_class αᵒᵈ αᵒᵈ (*) (≤) := ⟨c.1.flip⟩ @[to_additive] instance contravariant_class_swap_mul_le [has_le α] [has_mul α] [c : contravariant_class α α (swap (*)) (≤)] : - contravariant_class (order_dual α) (order_dual α) (swap (*)) (≤) := + contravariant_class αᵒᵈ αᵒᵈ (swap (*)) (≤) := ⟨c.1.flip⟩ @[to_additive] instance covariant_class_swap_mul_le [has_le α] [has_mul α] [c : covariant_class α α (swap (*)) (≤)] : - covariant_class (order_dual α) (order_dual α) (swap (*)) (≤) := + covariant_class αᵒᵈ αᵒᵈ (swap (*)) (≤) := ⟨c.1.flip⟩ @[to_additive] instance contravariant_class_mul_lt [has_lt α] [has_mul α] [c : contravariant_class α α (*) (<)] : - contravariant_class (order_dual α) (order_dual α) (*) (<) := + contravariant_class αᵒᵈ αᵒᵈ (*) (<) := ⟨c.1.flip⟩ @[to_additive] instance covariant_class_mul_lt [has_lt α] [has_mul α] [c : covariant_class α α (*) (<)] : - covariant_class (order_dual α) (order_dual α) (*) (<) := + covariant_class αᵒᵈ αᵒᵈ (*) (<) := ⟨c.1.flip⟩ @[to_additive] instance contravariant_class_swap_mul_lt [has_lt α] [has_mul α] [c : contravariant_class α α (swap (*)) (<)] : - contravariant_class (order_dual α) (order_dual α) (swap (*)) (<) := + contravariant_class αᵒᵈ αᵒᵈ (swap (*)) (<) := ⟨c.1.flip⟩ @[to_additive] instance covariant_class_swap_mul_lt [has_lt α] [has_mul α] [c : covariant_class α α (swap (*)) (<)] : - covariant_class (order_dual α) (order_dual α) (swap (*)) (<) := + covariant_class αᵒᵈ αᵒᵈ (swap (*)) (<) := ⟨c.1.flip⟩ @[to_additive] -instance [ordered_comm_monoid α] : ordered_comm_monoid (order_dual α) := +instance [ordered_comm_monoid α] : ordered_comm_monoid αᵒᵈ := { mul_le_mul_left := λ a b h c, mul_le_mul_left' h c, .. order_dual.partial_order α, .. order_dual.comm_monoid } @[to_additive ordered_cancel_add_comm_monoid.to_contravariant_class] instance ordered_cancel_comm_monoid.to_contravariant_class [ordered_cancel_comm_monoid α] : - contravariant_class (order_dual α) (order_dual α) has_mul.mul has_le.le := + contravariant_class αᵒᵈ αᵒᵈ has_mul.mul has_le.le := { elim := λ a b c bc, (ordered_cancel_comm_monoid.le_of_mul_le_mul_left a c b (dual_le.mp bc)) } @[to_additive] -instance [ordered_cancel_comm_monoid α] : ordered_cancel_comm_monoid (order_dual α) := +instance [ordered_cancel_comm_monoid α] : ordered_cancel_comm_monoid αᵒᵈ := { le_of_mul_le_mul_left := λ a b c : α, le_of_mul_le_mul_left', .. order_dual.ordered_comm_monoid, .. order_dual.cancel_comm_monoid } @[to_additive] instance [linear_ordered_cancel_comm_monoid α] : - linear_ordered_cancel_comm_monoid (order_dual α) := + linear_ordered_cancel_comm_monoid αᵒᵈ := { .. order_dual.linear_order α, .. order_dual.ordered_cancel_comm_monoid } @[to_additive] instance [linear_ordered_comm_monoid α] : - linear_ordered_comm_monoid (order_dual α) := + linear_ordered_comm_monoid αᵒᵈ := { .. order_dual.linear_order α, .. order_dual.ordered_comm_monoid } @@ -1089,7 +1089,7 @@ lemma coe_bit1 [has_one α] {a : α} : ((bit1 a : α) : with_bot α) = bit1 a := lemma add_ne_bot : a + b ≠ ⊥ ↔ a ≠ ⊥ ∧ b ≠ ⊥ := with_top.add_ne_top lemma bot_lt_add [partial_order α] {a b : with_bot α} : ⊥ < a + b ↔ ⊥ < a ∧ ⊥ < b := -@with_top.add_lt_top (order_dual α) _ _ _ _ +@with_top.add_lt_top αᵒᵈ _ _ _ _ lemma add_eq_coe : a + b = x ↔ ∃ (a' b' : α), ↑a' = a ∧ ↑b' = b ∧ a' + b' = x := with_top.add_eq_coe @@ -1100,35 +1100,35 @@ variables [preorder α] instance covariant_class_add_le [covariant_class α α (+) (≤)] : covariant_class (with_bot α) (with_bot α) (+) (≤) := -@order_dual.covariant_class_add_le (with_top $ order_dual α) _ _ _ +@order_dual.covariant_class_add_le (with_top αᵒᵈ) _ _ _ instance covariant_class_swap_add_le [covariant_class α α (swap (+)) (≤)] : covariant_class (with_bot α) (with_bot α) (swap (+)) (≤) := -@order_dual.covariant_class_swap_add_le (with_top $ order_dual α) _ _ _ +@order_dual.covariant_class_swap_add_le (with_top αᵒᵈ) _ _ _ instance contravariant_class_add_lt [contravariant_class α α (+) (<)] : contravariant_class (with_bot α) (with_bot α) (+) (<) := -@order_dual.contravariant_class_add_lt (with_top $ order_dual α) _ _ _ +@order_dual.contravariant_class_add_lt (with_top αᵒᵈ) _ _ _ instance contravariant_class_swap_add_lt [contravariant_class α α (swap (+)) (<)] : contravariant_class (with_bot α) (with_bot α) (swap (+)) (<) := -@order_dual.contravariant_class_swap_add_lt (with_top $ order_dual α) _ _ _ +@order_dual.contravariant_class_swap_add_lt (with_top αᵒᵈ) _ _ _ protected lemma le_of_add_le_add_left [contravariant_class α α (+) (≤)] (ha : a ≠ ⊥) (h : a + b ≤ a + c) : b ≤ c := -@with_top.le_of_add_le_add_left (order_dual α) _ _ _ _ _ _ ha h +@with_top.le_of_add_le_add_left αᵒᵈ _ _ _ _ _ _ ha h protected lemma le_of_add_le_add_right [contravariant_class α α (swap (+)) (≤)] (ha : a ≠ ⊥) (h : b + a ≤ c + a) : b ≤ c := -@with_top.le_of_add_le_add_right (order_dual α) _ _ _ _ _ _ ha h +@with_top.le_of_add_le_add_right αᵒᵈ _ _ _ _ _ _ ha h protected lemma add_lt_add_left [covariant_class α α (+) (<)] (ha : a ≠ ⊥) (h : b < c) : a + b < a + c := -@with_top.add_lt_add_left (order_dual α) _ _ _ _ _ _ ha h +@with_top.add_lt_add_left αᵒᵈ _ _ _ _ _ _ ha h protected lemma add_lt_add_right [covariant_class α α (swap (+)) (<)] (ha : a ≠ ⊥) (h : b < c) : b + a < c + a := -@with_top.add_lt_add_right (order_dual α) _ _ _ _ _ _ ha h +@with_top.add_lt_add_right αᵒᵈ _ _ _ _ _ _ ha h protected lemma add_le_add_iff_left [covariant_class α α (+) (≤)] [contravariant_class α α (+) (≤)] (ha : a ≠ ⊥) : a + b ≤ a + c ↔ b ≤ c := @@ -1148,11 +1148,11 @@ protected lemma add_lt_add_iff_right [covariant_class α α (swap (+)) (<)] protected lemma add_lt_add_of_le_of_lt [covariant_class α α (+) (<)] [covariant_class α α (swap (+)) (≤)] (hb : b ≠ ⊥) (hab : a ≤ b) (hcd : c < d) : a + c < b + d := -@with_top.add_lt_add_of_le_of_lt (order_dual α) _ _ _ _ _ _ _ _ hb hab hcd +@with_top.add_lt_add_of_le_of_lt αᵒᵈ _ _ _ _ _ _ _ _ hb hab hcd protected lemma add_lt_add_of_lt_of_le [covariant_class α α (+) (≤)] [covariant_class α α (swap (+)) (<)] (hd : d ≠ ⊥) (hab : a < b) (hcd : c ≤ d) : a + c < b + d := -@with_top.add_lt_add_of_lt_of_le (order_dual α) _ _ _ _ _ _ _ _ hd hab hcd +@with_top.add_lt_add_of_lt_of_le αᵒᵈ _ _ _ _ _ _ _ _ hd hab hcd end has_add end with_bot diff --git a/src/algebra/order/ring.lean b/src/algebra/order/ring.lean index fef4937123adf..5e964988d52e9 100644 --- a/src/algebra/order/ring.lean +++ b/src/algebra/order/ring.lean @@ -95,20 +95,20 @@ namespace order_dual /-! Note that `order_dual` does not satisfy any of the ordered ring typeclasses due to the `zero_le_one` field. -/ -instance [h : distrib α] : distrib (order_dual α) := h -instance [has_mul α] [h : has_distrib_neg α] : has_distrib_neg (order_dual α) := h -instance [h : non_unital_non_assoc_semiring α] : non_unital_non_assoc_semiring (order_dual α) := h -instance [h : non_unital_semiring α] : non_unital_semiring (order_dual α) := h -instance [h : non_assoc_semiring α] : non_assoc_semiring (order_dual α) := h -instance [h : semiring α] : semiring (order_dual α) := h -instance [h : non_unital_comm_semiring α] : non_unital_comm_semiring (order_dual α) := h -instance [h : comm_semiring α] : comm_semiring (order_dual α) := h -instance [h : non_unital_non_assoc_ring α] : non_unital_non_assoc_ring (order_dual α) := h -instance [h : non_unital_ring α] : non_unital_ring (order_dual α) := h -instance [h : non_assoc_ring α] : non_assoc_ring (order_dual α) := h -instance [h : ring α] : ring (order_dual α) := h -instance [h : non_unital_comm_ring α] : non_unital_comm_ring (order_dual α) := h -instance [h : comm_ring α] : comm_ring (order_dual α) := h +instance [h : distrib α] : distrib αᵒᵈ := h +instance [has_mul α] [h : has_distrib_neg α] : has_distrib_neg αᵒᵈ := h +instance [h : non_unital_non_assoc_semiring α] : non_unital_non_assoc_semiring αᵒᵈ := h +instance [h : non_unital_semiring α] : non_unital_semiring αᵒᵈ := h +instance [h : non_assoc_semiring α] : non_assoc_semiring αᵒᵈ := h +instance [h : semiring α] : semiring αᵒᵈ := h +instance [h : non_unital_comm_semiring α] : non_unital_comm_semiring αᵒᵈ := h +instance [h : comm_semiring α] : comm_semiring αᵒᵈ := h +instance [h : non_unital_non_assoc_ring α] : non_unital_non_assoc_ring αᵒᵈ := h +instance [h : non_unital_ring α] : non_unital_ring αᵒᵈ := h +instance [h : non_assoc_ring α] : non_assoc_ring αᵒᵈ := h +instance [h : ring α] : ring αᵒᵈ := h +instance [h : non_unital_comm_ring α] : non_unital_comm_ring αᵒᵈ := h +instance [h : comm_ring α] : comm_ring αᵒᵈ := h end order_dual diff --git a/src/algebra/order/smul.lean b/src/algebra/order/smul.lean index 85bc3a0a26666..f7f874104681f 100644 --- a/src/algebra/order/smul.lean +++ b/src/algebra/order/smul.lean @@ -49,40 +49,36 @@ namespace order_dual variables {R M : Type*} -instance [has_scalar R M] : has_scalar R (order_dual M) := -{ smul := λ k x, order_dual.rec (λ x', (k • x' : M)) x } +instance [has_scalar R M] : has_scalar R Mᵒᵈ := ⟨λ k x, order_dual.rec (λ x', (k • x' : M)) x⟩ -instance [has_zero R] [add_zero_class M] [h : smul_with_zero R M] : - smul_with_zero R (order_dual M) := +instance [has_zero R] [add_zero_class M] [h : smul_with_zero R M] : smul_with_zero R Mᵒᵈ := { zero_smul := λ m, order_dual.rec (zero_smul _) m, smul_zero := λ r, order_dual.rec (smul_zero' _) r, ..order_dual.has_scalar } -instance [monoid R] [mul_action R M] : mul_action R (order_dual M) := +instance [monoid R] [mul_action R M] : mul_action R Mᵒᵈ := { one_smul := λ m, order_dual.rec (one_smul _) m, mul_smul := λ r, order_dual.rec mul_smul r, ..order_dual.has_scalar } instance [monoid_with_zero R] [add_monoid M] [mul_action_with_zero R M] : - mul_action_with_zero R (order_dual M) := + mul_action_with_zero R Mᵒᵈ := { ..order_dual.mul_action, ..order_dual.smul_with_zero } instance [monoid_with_zero R] [add_monoid M] [distrib_mul_action R M] : - distrib_mul_action R (order_dual M) := + distrib_mul_action R Mᵒᵈ := { smul_add := λ k a, order_dual.rec (λ a' b, order_dual.rec (smul_add _ _) b) a, smul_zero := λ r, order_dual.rec smul_zero r } instance [ordered_semiring R] [ordered_add_comm_monoid M] [smul_with_zero R M] [ordered_smul R M] : - ordered_smul R (order_dual M) := + ordered_smul R Mᵒᵈ := { smul_lt_smul_of_pos := λ a b, @ordered_smul.smul_lt_smul_of_pos R M _ _ _ _ b a, lt_of_smul_lt_smul_of_pos := λ a b, @ordered_smul.lt_of_smul_lt_smul_of_pos R M _ _ _ _ b a } @[simp] lemma to_dual_smul [has_scalar R M] {c : R} {a : M} : to_dual (c • a) = c • to_dual a := rfl - -@[simp] lemma of_dual_smul [has_scalar R M] {c : R} {a : order_dual M} : - of_dual (c • a) = c • of_dual a := +@[simp] lemma of_dual_smul [has_scalar R M] {c : R} {a : Mᵒᵈ} : of_dual (c • a) = c • of_dual a := rfl end order_dual @@ -109,7 +105,7 @@ calc (0 : M) = c • (0 : M) : (smul_zero' M c).symm ... ≤ c • a : smul_le_smul_of_nonneg ha hc lemma smul_nonpos_of_nonneg_of_nonpos (hc : 0 ≤ c) (ha : a ≤ 0) : c • a ≤ 0 := -@smul_nonneg R (order_dual M) _ _ _ _ _ _ hc ha +@smul_nonneg R Mᵒᵈ _ _ _ _ _ _ hc ha lemma eq_of_smul_eq_smul_of_pos_of_le (h₁ : c • a = c • b) (hc : 0 < c) (hle : a ≤ b) : a = b := diff --git a/src/algebra/order/with_zero.lean b/src/algebra/order/with_zero.lean index 8a199a22621e6..33eba665df98d 100644 --- a/src/algebra/order/with_zero.lean +++ b/src/algebra/order/with_zero.lean @@ -34,7 +34,7 @@ variables {α : Type*} variables {a b c d x y z : α} instance [linear_ordered_add_comm_monoid_with_top α] : - linear_ordered_comm_monoid_with_zero (multiplicative (order_dual α)) := + linear_ordered_comm_monoid_with_zero (multiplicative αᵒᵈ) := { zero := multiplicative.of_add (⊤ : α), zero_mul := top_add, mul_zero := add_top, @@ -43,7 +43,7 @@ instance [linear_ordered_add_comm_monoid_with_top α] : ..multiplicative.linear_order } instance [linear_ordered_add_comm_group_with_top α] : - linear_ordered_comm_group_with_zero (multiplicative (order_dual α)) := + linear_ordered_comm_group_with_zero (multiplicative αᵒᵈ) := { inv_zero := linear_ordered_add_comm_group_with_top.neg_top, mul_inv_cancel := linear_ordered_add_comm_group_with_top.add_neg_cancel, ..multiplicative.div_inv_monoid, @@ -196,7 +196,7 @@ lemma ne_zero_of_lt (h : b < a) : a ≠ 0 := lemma pow_pos_iff [no_zero_divisors α] {n : ℕ} (hn : 0 < n) : 0 < a ^ n ↔ 0 < a := by simp_rw [zero_lt_iff, pow_ne_zero_iff hn] -instance : linear_ordered_add_comm_monoid_with_top (additive (order_dual α)) := +instance : linear_ordered_add_comm_monoid_with_top (additive αᵒᵈ) := { top := (0 : α), top_add' := λ a, (zero_mul a : (0 : α) * a = 0), le_top := λ _, zero_le', @@ -294,7 +294,7 @@ by rw [div_eq_mul_inv, le_mul_inv_iff₀ hc] lemma div_le_iff₀ (hc : c ≠ 0) : a/c ≤ b ↔ a ≤ b*c := by rw [div_eq_mul_inv, mul_inv_le_iff₀ hc] -instance : linear_ordered_add_comm_group_with_top (additive (order_dual α)) := +instance : linear_ordered_add_comm_group_with_top (additive αᵒᵈ) := { neg_top := inv_zero, add_neg_cancel := λ a ha, mul_inv_cancel ha, ..additive.sub_neg_monoid, diff --git a/src/algebra/support.lean b/src/algebra/support.lean index d900375d0d67d..9b9b9dc8198b0 100644 --- a/src/algebra/support.lean +++ b/src/algebra/support.lean @@ -127,7 +127,7 @@ end @[to_additive] lemma mul_support_infi [conditionally_complete_lattice M] [nonempty ι] (f : ι → α → M) : mul_support (λ x, ⨅ i, f i x) ⊆ ⋃ i, mul_support (f i) := -@mul_support_supr _ (order_dual M) ι ⟨(1:M)⟩ _ _ f +@mul_support_supr _ Mᵒᵈ ι ⟨(1:M)⟩ _ _ f @[to_additive] lemma mul_support_comp_subset {g : M → N} (hg : g 1 = 1) (f : α → M) : mul_support (g ∘ f) ⊆ mul_support f := diff --git a/src/algebraic_geometry/prime_spectrum/basic.lean b/src/algebraic_geometry/prime_spectrum/basic.lean index 1898859ae7fe1..d45eb115408ec 100644 --- a/src/algebraic_geometry/prime_spectrum/basic.lean +++ b/src/algebraic_geometry/prime_spectrum/basic.lean @@ -161,14 +161,12 @@ section gc variable (R) /-- `zero_locus` and `vanishing_ideal` form a galois connection. -/ -lemma gc : @galois_connection - (ideal R) (order_dual (set (prime_spectrum R))) _ _ +lemma gc : @galois_connection (ideal R) (set (prime_spectrum R))ᵒᵈ _ _ (λ I, zero_locus I) (λ t, vanishing_ideal t) := λ I t, subset_zero_locus_iff_le_vanishing_ideal t I /-- `zero_locus` and `vanishing_ideal` form a galois connection. -/ -lemma gc_set : @galois_connection - (set R) (order_dual (set (prime_spectrum R))) _ _ +lemma gc_set : @galois_connection (set R) (set (prime_spectrum R))ᵒᵈ _ _ (λ s, zero_locus s) (λ t, vanishing_ideal t) := have ideal_gc : galois_connection (ideal.span) coe := (submodule.gi R R).gc, by simpa [zero_locus_span, function.comp] using ideal_gc.compose (gc R) diff --git a/src/algebraic_geometry/projective_spectrum/topology.lean b/src/algebraic_geometry/projective_spectrum/topology.lean index 31c1aceae6351..68e82cd8de7e5 100644 --- a/src/algebraic_geometry/projective_spectrum/topology.lean +++ b/src/algebraic_geometry/projective_spectrum/topology.lean @@ -130,19 +130,19 @@ lemma subset_zero_locus_iff_le_vanishing_ideal (t : set (projective_spectrum variable (𝒜) /-- `zero_locus` and `vanishing_ideal` form a galois connection. -/ lemma gc_ideal : @galois_connection - (ideal A) (order_dual (set (projective_spectrum 𝒜))) _ _ + (ideal A) (set (projective_spectrum 𝒜))ᵒᵈ _ _ (λ I, zero_locus 𝒜 I) (λ t, (vanishing_ideal t).to_ideal) := λ I t, subset_zero_locus_iff_le_vanishing_ideal t I /-- `zero_locus` and `vanishing_ideal` form a galois connection. -/ lemma gc_set : @galois_connection - (set A) (order_dual (set (projective_spectrum 𝒜))) _ _ + (set A) (set (projective_spectrum 𝒜))ᵒᵈ _ _ (λ s, zero_locus 𝒜 s) (λ t, vanishing_ideal t) := have ideal_gc : galois_connection (ideal.span) coe := (submodule.gi A _).gc, by simpa [zero_locus_span, function.comp] using galois_connection.compose ideal_gc (gc_ideal 𝒜) lemma gc_homogeneous_ideal : @galois_connection - (homogeneous_ideal 𝒜) (order_dual (set (projective_spectrum 𝒜))) _ _ + (homogeneous_ideal 𝒜) (set (projective_spectrum 𝒜))ᵒᵈ _ _ (λ I, zero_locus 𝒜 I) (λ t, (vanishing_ideal t)) := λ I t, by simpa [show I.to_ideal ≤ (vanishing_ideal t).to_ideal ↔ I ≤ (vanishing_ideal t), from iff.rfl] using subset_zero_locus_iff_le_vanishing_ideal t I.to_ideal diff --git a/src/analysis/box_integral/partition/filter.lean b/src/analysis/box_integral/partition/filter.lean index 2c876aa786f47..b73ff899457a2 100644 --- a/src/analysis/box_integral/partition/filter.lean +++ b/src/analysis/box_integral/partition/filter.lean @@ -193,7 +193,7 @@ variables {l l₁ l₂ : integration_params} namespace integration_params /-- Auxiliary equivalence with a product type used to lift an order. -/ -def equiv_prod : integration_params ≃ bool × order_dual bool × order_dual bool := +def equiv_prod : integration_params ≃ bool × boolᵒᵈ × boolᵒᵈ := { to_fun := λ l, ⟨l.1, order_dual.to_dual l.2, order_dual.to_dual l.3⟩, inv_fun := λ l, ⟨l.1, order_dual.of_dual l.2.1, order_dual.of_dual l.2.2⟩, left_inv := λ ⟨a, b, c⟩, rfl, @@ -203,7 +203,7 @@ instance : partial_order integration_params := partial_order.lift equiv_prod equiv_prod.injective /-- Auxiliary `order_iso` with a product type used to lift a `bounded_order` structure. -/ -def iso_prod : integration_params ≃o bool × order_dual bool × order_dual bool := +def iso_prod : integration_params ≃o bool × boolᵒᵈ × boolᵒᵈ := ⟨equiv_prod, λ ⟨x, y, z⟩, iff.rfl⟩ instance : bounded_order integration_params := diff --git a/src/analysis/convex/basic.lean b/src/analysis/convex/basic.lean index c3b4b10c672a0..162b16ef7cba2 100644 --- a/src/analysis/convex/basic.lean +++ b/src/analysis/convex/basic.lean @@ -696,8 +696,7 @@ calc : add_le_add (smul_le_smul_of_nonneg hx ha) (smul_le_smul_of_nonneg hy hb) ... = r : convex.combo_self hab _ -lemma convex_Ici (r : β) : convex 𝕜 (Ici r) := -@convex_Iic 𝕜 (order_dual β) _ _ _ _ r +lemma convex_Ici (r : β) : convex 𝕜 (Ici r) := @convex_Iic 𝕜 βᵒᵈ _ _ _ _ r lemma convex_Icc (r s : β) : convex 𝕜 (Icc r s) := Ici_inter_Iic.subst ((convex_Ici r).inter $ convex_Iic s) @@ -736,8 +735,7 @@ begin ... = r : convex.combo_self hab _ end -lemma convex_Ioi (r : β) : convex 𝕜 (Ioi r) := -@convex_Iio 𝕜 (order_dual β) _ _ _ _ r +lemma convex_Ioi (r : β) : convex 𝕜 (Ioi r) := @convex_Iio 𝕜 βᵒᵈ _ _ _ _ r lemma convex_Ioo (r s : β) : convex 𝕜 (Ioo r s) := Ioi_inter_Iio.subst ((convex_Ioi r).inter $ convex_Iio s) @@ -786,27 +784,27 @@ lemma monotone_on.convex_lt (hf : monotone_on f s) (hs : convex 𝕜 s) (r : β) lemma monotone_on.convex_ge (hf : monotone_on f s) (hs : convex 𝕜 s) (r : β) : convex 𝕜 {x ∈ s | r ≤ f x} := -@monotone_on.convex_le 𝕜 (order_dual E) (order_dual β) _ _ _ _ _ _ _ hf.dual hs r +@monotone_on.convex_le 𝕜 Eᵒᵈ βᵒᵈ _ _ _ _ _ _ _ hf.dual hs r lemma monotone_on.convex_gt (hf : monotone_on f s) (hs : convex 𝕜 s) (r : β) : convex 𝕜 {x ∈ s | r < f x} := -@monotone_on.convex_lt 𝕜 (order_dual E) (order_dual β) _ _ _ _ _ _ _ hf.dual hs r +@monotone_on.convex_lt 𝕜 Eᵒᵈ βᵒᵈ _ _ _ _ _ _ _ hf.dual hs r lemma antitone_on.convex_le (hf : antitone_on f s) (hs : convex 𝕜 s) (r : β) : convex 𝕜 {x ∈ s | f x ≤ r} := -@monotone_on.convex_ge 𝕜 E (order_dual β) _ _ _ _ _ _ _ hf hs r +@monotone_on.convex_ge 𝕜 E βᵒᵈ _ _ _ _ _ _ _ hf hs r lemma antitone_on.convex_lt (hf : antitone_on f s) (hs : convex 𝕜 s) (r : β) : convex 𝕜 {x ∈ s | f x < r} := -@monotone_on.convex_gt 𝕜 E (order_dual β) _ _ _ _ _ _ _ hf hs r +@monotone_on.convex_gt 𝕜 E βᵒᵈ _ _ _ _ _ _ _ hf hs r lemma antitone_on.convex_ge (hf : antitone_on f s) (hs : convex 𝕜 s) (r : β) : convex 𝕜 {x ∈ s | r ≤ f x} := -@monotone_on.convex_le 𝕜 E (order_dual β) _ _ _ _ _ _ _ hf hs r +@monotone_on.convex_le 𝕜 E βᵒᵈ _ _ _ _ _ _ _ hf hs r lemma antitone_on.convex_gt (hf : antitone_on f s) (hs : convex 𝕜 s) (r : β) : convex 𝕜 {x ∈ s | r < f x} := -@monotone_on.convex_lt 𝕜 E (order_dual β) _ _ _ _ _ _ _ hf hs r +@monotone_on.convex_lt 𝕜 E βᵒᵈ _ _ _ _ _ _ _ hf hs r lemma monotone.convex_le (hf : monotone f) (r : β) : convex 𝕜 {x | f x ≤ r} := diff --git a/src/analysis/convex/extrema.lean b/src/analysis/convex/extrema.lean index 5bc221d88a1c9..340be610d6774 100644 --- a/src/analysis/convex/extrema.lean +++ b/src/analysis/convex/extrema.lean @@ -74,8 +74,7 @@ end lemma is_max_on.of_is_local_max_on_of_concave_on {f : E → β} {a : E} (a_in_s : a ∈ s) (h_localmax: is_local_max_on f s a) (h_conc : concave_on ℝ s f) : is_max_on f s a := -@is_min_on.of_is_local_min_on_of_convex_on - _ (order_dual β) _ _ _ _ _ _ _ _ s f a a_in_s h_localmax h_conc +@is_min_on.of_is_local_min_on_of_convex_on _ βᵒᵈ _ _ _ _ _ _ _ _ s f a a_in_s h_localmax h_conc /-- A local minimum of a convex function is a global minimum. -/ lemma is_min_on.of_is_local_min_of_convex_univ {f : E → β} {a : E} @@ -86,4 +85,4 @@ lemma is_min_on.of_is_local_min_of_convex_univ {f : E → β} {a : E} /-- A local maximum of a concave function is a global maximum. -/ lemma is_max_on.of_is_local_max_of_convex_univ {f : E → β} {a : E} (h_local_max : is_local_max f a) (h_conc : concave_on ℝ univ f) : ∀ x, f x ≤ f a := -@is_min_on.of_is_local_min_of_convex_univ _ (order_dual β) _ _ _ _ _ _ _ _ f a h_local_max h_conc +@is_min_on.of_is_local_min_of_convex_univ _ βᵒᵈ _ _ _ _ _ _ _ _ f a h_local_max h_conc diff --git a/src/analysis/convex/function.lean b/src/analysis/convex/function.lean index bcc1464ef2596..94cd592ec2d38 100644 --- a/src/analysis/convex/function.lean +++ b/src/analysis/convex/function.lean @@ -131,7 +131,7 @@ lemma convex_on_const (c : β) (hs : convex 𝕜 s) : convex_on 𝕜 s (λ x:E, ⟨hs, λ x y _ _ a b _ _ hab, (convex.combo_self hab c).ge⟩ lemma concave_on_const (c : β) (hs : convex 𝕜 s) : concave_on 𝕜 s (λ x:E, c) := -@convex_on_const _ _ (order_dual β) _ _ _ _ _ _ c hs +@convex_on_const _ _ βᵒᵈ _ _ _ _ _ _ c hs lemma convex_on_of_convex_epigraph (h : convex 𝕜 {p : E × β | p.1 ∈ s ∧ f p.1 ≤ p.2}) : convex_on 𝕜 s f := @@ -140,7 +140,7 @@ lemma convex_on_of_convex_epigraph (h : convex 𝕜 {p : E × β | p.1 ∈ s ∧ lemma concave_on_of_convex_hypograph (h : convex 𝕜 {p : E × β | p.1 ∈ s ∧ p.2 ≤ f p.1}) : concave_on 𝕜 s f := -@convex_on_of_convex_epigraph 𝕜 E (order_dual β) _ _ _ _ _ _ _ h +@convex_on_of_convex_epigraph 𝕜 E βᵒᵈ _ _ _ _ _ _ _ h end module @@ -180,7 +180,7 @@ lemma convex_on_iff_convex_epigraph : lemma concave_on_iff_convex_hypograph : concave_on 𝕜 s f ↔ convex 𝕜 {p : E × β | p.1 ∈ s ∧ p.2 ≤ f p.1} := -@convex_on_iff_convex_epigraph 𝕜 E (order_dual β) _ _ _ _ _ _ _ f +@convex_on_iff_convex_epigraph 𝕜 E βᵒᵈ _ _ _ _ _ _ _ f end ordered_smul @@ -234,7 +234,7 @@ lemma concave_on_iff_forall_pos {s : set E} {f : E → β} : concave_on 𝕜 s f ↔ convex 𝕜 s ∧ ∀ ⦃x y : E⦄, x ∈ s → y ∈ s → ∀ ⦃a b : 𝕜⦄, 0 < a → 0 < b → a + b = 1 → a • f x + b • f y ≤ f (a • x + b • y) := -@convex_on_iff_forall_pos 𝕜 E (order_dual β) _ _ _ _ _ _ _ +@convex_on_iff_forall_pos 𝕜 E βᵒᵈ _ _ _ _ _ _ _ lemma convex_on_iff_pairwise_pos {s : set E} {f : E → β} : convex_on 𝕜 s f ↔ convex 𝕜 s ∧ @@ -253,7 +253,7 @@ lemma concave_on_iff_pairwise_pos {s : set E} {f : E → β} : concave_on 𝕜 s f ↔ convex 𝕜 s ∧ s.pairwise (λ x y, ∀ ⦃a b : 𝕜⦄, 0 < a → 0 < b → a + b = 1 → a • f x + b • f y ≤ f (a • x + b • y)) := -@convex_on_iff_pairwise_pos 𝕜 E (order_dual β) _ _ _ _ _ _ _ +@convex_on_iff_pairwise_pos 𝕜 E βᵒᵈ _ _ _ _ _ _ _ /-- A linear map is convex. -/ lemma linear_map.convex_on (f : E →ₗ[𝕜] β) {s : set E} (hs : convex 𝕜 s) : convex_on 𝕜 s f := @@ -314,7 +314,7 @@ main use case is `E = ℝ` however one can apply it, e.g., to `ℝ^n` with lexic lemma linear_order.concave_on_of_lt (hs : convex 𝕜 s) (hf : ∀ ⦃x y : E⦄, x ∈ s → y ∈ s → x < y → ∀ ⦃a b : 𝕜⦄, 0 < a → 0 < b → a + b = 1 → a • f x + b • f y ≤ f (a • x + b • y)) : concave_on 𝕜 s f := -@linear_order.convex_on_of_lt _ _ (order_dual β) _ _ _ _ _ _ s f hs hf +@linear_order.convex_on_of_lt _ _ βᵒᵈ _ _ _ _ _ _ s f hs hf /-- For a function on a convex set in a linearly ordered space (where the order and the algebraic structures aren't necessarily compatible), in order to prove that it is convex, it suffices to @@ -337,7 +337,7 @@ main use case is `E = 𝕜` however one can apply it, e.g., to `𝕜^n` with lex lemma linear_order.strict_concave_on_of_lt (hs : convex 𝕜 s) (hf : ∀ ⦃x y : E⦄, x ∈ s → y ∈ s → x < y → ∀ ⦃a b : 𝕜⦄, 0 < a → 0 < b → a + b = 1 → a • f x + b • f y < f (a • x + b • y)) : strict_concave_on 𝕜 s f := -@linear_order.strict_convex_on_of_lt _ _ (order_dual β) _ _ _ _ _ _ _ _ hs hf +@linear_order.strict_convex_on_of_lt _ _ βᵒᵈ _ _ _ _ _ _ _ _ hs hf end linear_order end module @@ -824,7 +824,7 @@ end⟩ lemma concave_on_iff_div {f : E → β} : concave_on 𝕜 s f ↔ convex 𝕜 s ∧ ∀ ⦃x y : E⦄, x ∈ s → y ∈ s → ∀ ⦃a b : 𝕜⦄, 0 ≤ a → 0 ≤ b → 0 < a + b → (a/(a+b)) • f x + (b/(a+b)) • f y ≤ f ((a/(a+b)) • x + (b/(a+b)) • y) := -@convex_on_iff_div _ _ (order_dual β) _ _ _ _ _ _ _ +@convex_on_iff_div _ _ βᵒᵈ _ _ _ _ _ _ _ lemma strict_convex_on_iff_div {f : E → β} : strict_convex_on 𝕜 s f ↔ convex 𝕜 s ∧ ∀ ⦃x y : E⦄, x ∈ s → y ∈ s → x ≠ y → ∀ ⦃a b : 𝕜⦄, 0 < a @@ -844,7 +844,7 @@ end⟩ lemma strict_concave_on_iff_div {f : E → β} : strict_concave_on 𝕜 s f ↔ convex 𝕜 s ∧ ∀ ⦃x y : E⦄, x ∈ s → y ∈ s → x ≠ y → ∀ ⦃a b : 𝕜⦄, 0 < a → 0 < b → (a/(a+b)) • f x + (b/(a+b)) • f y < f ((a/(a+b)) • x + (b/(a+b)) • y) := -@strict_convex_on_iff_div _ _ (order_dual β) _ _ _ _ _ _ _ +@strict_convex_on_iff_div _ _ βᵒᵈ _ _ _ _ _ _ _ end has_scalar end ordered_add_comm_monoid diff --git a/src/analysis/convex/jensen.lean b/src/analysis/convex/jensen.lean index a39c01298cb18..a548eb26f0c6b 100644 --- a/src/analysis/convex/jensen.lean +++ b/src/analysis/convex/jensen.lean @@ -52,7 +52,7 @@ end lemma concave_on.le_map_center_mass (hf : concave_on 𝕜 s f) (h₀ : ∀ i ∈ t, 0 ≤ w i) (h₁ : 0 < ∑ i in t, w i) (hmem : ∀ i ∈ t, p i ∈ s) : t.center_mass w (f ∘ p) ≤ f (t.center_mass w p) := -@convex_on.map_center_mass_le 𝕜 E (order_dual β) _ _ _ _ _ _ _ _ _ _ _ _ hf h₀ h₁ hmem +@convex_on.map_center_mass_le 𝕜 E βᵒᵈ _ _ _ _ _ _ _ _ _ _ _ _ hf h₀ h₁ hmem /-- Convex **Jensen's inequality**, `finset.sum` version. -/ lemma convex_on.map_sum_le (hf : convex_on 𝕜 s f) (h₀ : ∀ i ∈ t, 0 ≤ w i) (h₁ : ∑ i in t, w i = 1) @@ -65,7 +65,7 @@ by simpa only [center_mass, h₁, inv_one, one_smul] lemma concave_on.le_map_sum (hf : concave_on 𝕜 s f) (h₀ : ∀ i ∈ t, 0 ≤ w i) (h₁ : ∑ i in t, w i = 1) (hmem : ∀ i ∈ t, p i ∈ s) : ∑ i in t, w i • f (p i) ≤ f (∑ i in t, w i • p i) := -@convex_on.map_sum_le 𝕜 E (order_dual β) _ _ _ _ _ _ _ _ _ _ _ _ hf h₀ h₁ hmem +@convex_on.map_sum_le 𝕜 E βᵒᵈ _ _ _ _ _ _ _ _ _ _ _ _ hf h₀ h₁ hmem end jensen @@ -100,7 +100,7 @@ end lemma concave_on.exists_le_of_center_mass (h : concave_on 𝕜 s f) (hw₀ : ∀ i ∈ t, 0 ≤ w i) (hw₁ : 0 < ∑ i in t, w i) (hp : ∀ i ∈ t, p i ∈ s) : ∃ i ∈ t, f (p i) ≤ f (t.center_mass w p) := -@convex_on.exists_ge_of_center_mass 𝕜 E (order_dual β) _ _ _ _ _ _ _ _ _ _ _ _ h hw₀ hw₁ hp +@convex_on.exists_ge_of_center_mass 𝕜 E βᵒᵈ _ _ _ _ _ _ _ _ _ _ _ _ h hw₀ hw₁ hp /-- Maximum principle for convex functions. If a function `f` is convex on the convex hull of `s`, then the eventual maximum of `f` on `convex_hull 𝕜 s` lies in `s`. -/ @@ -118,6 +118,6 @@ end then the eventual minimum of `f` on `convex_hull 𝕜 s` lies in `s`. -/ lemma concave_on.exists_le_of_mem_convex_hull (hf : concave_on 𝕜 (convex_hull 𝕜 s) f) {x} (hx : x ∈ convex_hull 𝕜 s) : ∃ y ∈ s, f y ≤ f x := -@convex_on.exists_ge_of_mem_convex_hull 𝕜 E (order_dual β) _ _ _ _ _ _ _ _ hf _ hx +@convex_on.exists_ge_of_mem_convex_hull 𝕜 E βᵒᵈ _ _ _ _ _ _ _ _ hf _ hx end maximum_principle diff --git a/src/analysis/convex/quasiconvex.lean b/src/analysis/convex/quasiconvex.lean index a61ed4d00ffc2..6fd19188f29d9 100644 --- a/src/analysis/convex/quasiconvex.lean +++ b/src/analysis/convex/quasiconvex.lean @@ -80,7 +80,7 @@ lemma convex.quasiconvex_on_of_convex_le (hs : convex 𝕜 s) (h : ∀ r, convex lemma convex.quasiconcave_on_of_convex_ge (hs : convex 𝕜 s) (h : ∀ r, convex 𝕜 {x | r ≤ f x}) : quasiconcave_on 𝕜 s f := -@convex.quasiconvex_on_of_convex_le 𝕜 E (order_dual β) _ _ _ _ _ _ hs h +@convex.quasiconvex_on_of_convex_le 𝕜 E βᵒᵈ _ _ _ _ _ _ hs h lemma quasiconvex_on.convex [is_directed β (≤)] (hf : quasiconvex_on 𝕜 s f) : convex 𝕜 s := λ x y hx hy a b ha hb hab, @@ -122,7 +122,7 @@ lemma quasiconcave_on_iff_min_le : quasiconcave_on 𝕜 s f ↔ convex 𝕜 s ∧ ∀ ⦃x y : E⦄, x ∈ s → y ∈ s → ∀ ⦃a b : 𝕜⦄, 0 ≤ a → 0 ≤ b → a + b = 1 → min (f x) (f y) ≤ f (a • x + b • y) := -@quasiconvex_on_iff_le_max 𝕜 E (order_dual β) _ _ _ _ _ _ +@quasiconvex_on_iff_le_max 𝕜 E βᵒᵈ _ _ _ _ _ _ lemma quasilinear_on_iff_mem_interval : quasilinear_on 𝕜 s f ↔ convex 𝕜 s ∧ diff --git a/src/analysis/convex/strict.lean b/src/analysis/convex/strict.lean index 935c86d37a588..fae867240a565 100644 --- a/src/analysis/convex/strict.lean +++ b/src/analysis/convex/strict.lean @@ -153,8 +153,7 @@ begin { exact add_lt_add (smul_lt_smul_of_pos hx ha) (smul_lt_smul_of_pos hy hb) } end -lemma strict_convex_Ici (r : β) : strict_convex 𝕜 (Ici r) := -@strict_convex_Iic 𝕜 (order_dual β) _ _ _ _ _ _ r +lemma strict_convex_Ici (r : β) : strict_convex 𝕜 (Ici r) := @strict_convex_Iic 𝕜 βᵒᵈ _ _ _ _ _ _ r lemma strict_convex_Icc (r s : β) : strict_convex 𝕜 (Icc r s) := (strict_convex_Ici r).inter $ strict_convex_Iic s diff --git a/src/analysis/inner_product_space/basic.lean b/src/analysis/inner_product_space/basic.lean index 3f3ef98cff338..7414b5e7a66f7 100644 --- a/src/analysis/inner_product_space/basic.lean +++ b/src/analysis/inner_product_space/basic.lean @@ -2203,7 +2203,7 @@ variables (𝕜 E) /-- `submodule.orthogonal` gives a `galois_connection` between `submodule 𝕜 E` and its `order_dual`. -/ lemma submodule.orthogonal_gc : - @galois_connection (submodule 𝕜 E) (order_dual $ submodule 𝕜 E) _ _ + @galois_connection (submodule 𝕜 E) (submodule 𝕜 E)ᵒᵈ _ _ submodule.orthogonal submodule.orthogonal := λ K₁ K₂, ⟨λ h v hv u hu, submodule.inner_left_of_mem_orthogonal hv (h hu), λ h v hv u hu, submodule.inner_left_of_mem_orthogonal hv (h hu)⟩ diff --git a/src/analysis/normed_space/lattice_ordered_group.lean b/src/analysis/normed_space/lattice_ordered_group.lean index b6046e7258ea8..447bccf3bda2c 100644 --- a/src/analysis/normed_space/lattice_ordered_group.lean +++ b/src/analysis/normed_space/lattice_ordered_group.lean @@ -63,7 +63,7 @@ instance normed_lattice_add_comm_group_to_ordered_add_comm_group {α : Type*} Let `α` be a normed group with a partial order. Then the order dual is also a normed group. -/ @[priority 100] -- see Note [lower instance priority] -instance {α : Type*} : Π [normed_group α], normed_group (order_dual α) := id +instance {α : Type*} : Π [normed_group α], normed_group αᵒᵈ := id variables {α : Type*} [normed_lattice_add_comm_group α] open lattice_ordered_comm_group @@ -84,7 +84,7 @@ Let `α` be a normed lattice ordered group, then the order dual is also a normed lattice ordered group. -/ @[priority 100] -- see Note [lower instance priority] -instance : normed_lattice_add_comm_group (order_dual α) := +instance : normed_lattice_add_comm_group αᵒᵈ := { add_le_add_left := begin intros a b h₁ c, rw ← order_dual.dual_le, @@ -150,7 +150,7 @@ end instance normed_lattice_add_comm_group_has_continuous_sup {α : Type*} [normed_lattice_add_comm_group α] : has_continuous_sup α := -order_dual.has_continuous_sup (order_dual α) +order_dual.has_continuous_sup αᵒᵈ /-- Let `α` be a normed lattice ordered group. Then `α` is a topological lattice in the norm topology. diff --git a/src/combinatorics/pigeonhole.lean b/src/combinatorics/pigeonhole.lean index 4b1a755f04ea4..7c3655380f457 100644 --- a/src/combinatorics/pigeonhole.lean +++ b/src/combinatorics/pigeonhole.lean @@ -124,7 +124,7 @@ than `b`. -/ lemma exists_sum_fiber_lt_of_maps_to_of_sum_lt_nsmul (hf : ∀ a ∈ s, f a ∈ t) (hb : (∑ x in s, w x) < t.card • b) : ∃ y ∈ t, (∑ x in s.filter (λ x, f x = y), w x) < b := -@exists_lt_sum_fiber_of_maps_to_of_nsmul_lt_sum α β (order_dual M) _ _ _ _ _ _ _ hf hb +@exists_lt_sum_fiber_of_maps_to_of_nsmul_lt_sum α β Mᵒᵈ _ _ _ _ _ _ _ hf hb /-- The pigeonhole principle for finitely many pigeons counted by weight, strict inequality version: if the total weight of a finite set of pigeons is greater than `n • b`, they are sorted into some @@ -147,7 +147,7 @@ is less than `b`. -/ lemma exists_sum_fiber_lt_of_sum_fiber_nonneg_of_sum_lt_nsmul (ht : ∀ y ∉ t, (0:M) ≤ ∑ x in s.filter (λ x, f x = y), w x) (hb : (∑ x in s, w x) < t.card • b) : ∃ y ∈ t, (∑ x in s.filter (λ x, f x = y), w x) < b := -@exists_lt_sum_fiber_of_sum_fiber_nonpos_of_nsmul_lt_sum α β (order_dual M) _ _ _ _ _ _ _ ht hb +@exists_lt_sum_fiber_of_sum_fiber_nonpos_of_nsmul_lt_sum α β Mᵒᵈ _ _ _ _ _ _ _ ht hb /-! #### Non-strict inequality versions @@ -169,7 +169,7 @@ this pigeonhole is less than or equal to `b`. -/ lemma exists_sum_fiber_le_of_maps_to_of_sum_le_nsmul (hf : ∀ a ∈ s, f a ∈ t) (ht : t.nonempty) (hb : (∑ x in s, w x) ≤ t.card • b) : ∃ y ∈ t, (∑ x in s.filter (λ x, f x = y), w x) ≤ b := -@exists_le_sum_fiber_of_maps_to_of_nsmul_le_sum α β (order_dual M) _ _ _ _ _ _ _ hf ht hb +@exists_le_sum_fiber_of_maps_to_of_nsmul_le_sum α β Mᵒᵈ _ _ _ _ _ _ _ hf ht hb /-- The pigeonhole principle for finitely many pigeons counted by weight, non-strict inequality version: if the total weight of a finite set of pigeons is greater than or equal to `n • b`, they @@ -194,7 +194,7 @@ lemma exists_sum_fiber_le_of_sum_fiber_nonneg_of_sum_le_nsmul (hf : ∀ y ∉ t, (0:M) ≤ ∑ x in s.filter (λ x, f x = y), w x) (ht : t.nonempty) (hb : (∑ x in s, w x) ≤ t.card • b) : ∃ y ∈ t, (∑ x in s.filter (λ x, f x = y), w x) ≤ b := -@exists_le_sum_fiber_of_sum_fiber_nonpos_of_nsmul_le_sum α β (order_dual M) _ _ _ _ _ _ _ hf ht hb +@exists_le_sum_fiber_of_sum_fiber_nonpos_of_nsmul_le_sum α β Mᵒᵈ _ _ _ _ _ _ _ hf ht hb end @@ -356,7 +356,7 @@ version: there is a pigeonhole with the total weight of pigeons in it less than the total number of pigeonholes times `b` is greater than the total weight of all pigeons. -/ lemma exists_sum_fiber_lt_of_sum_lt_nsmul (hb : (∑ x, w x) < card β • b) : ∃ y, (∑ x in univ.filter (λ x, f x = y), w x) < b := -@exists_lt_sum_fiber_of_nsmul_lt_sum α β (order_dual M) _ _ _ _ _ _ _ hb +@exists_lt_sum_fiber_of_nsmul_lt_sum α β Mᵒᵈ _ _ _ _ _ _ _ hb /-- The pigeonhole principle for finitely many pigeons of different weights, non-strict inequality version: there is a pigeonhole with the total weight of pigeons in it less than or equal to `b` @@ -364,7 +364,7 @@ provided that the total number of pigeonholes times `b` is greater than or equal of all pigeons. -/ lemma exists_sum_fiber_le_of_sum_le_nsmul [nonempty β] (hb : (∑ x, w x) ≤ card β • b) : ∃ y, (∑ x in univ.filter (λ x, f x = y), w x) ≤ b := -@exists_le_sum_fiber_of_nsmul_le_sum α β (order_dual M) _ _ _ _ _ _ _ _ hb +@exists_le_sum_fiber_of_nsmul_le_sum α β Mᵒᵈ _ _ _ _ _ _ _ _ hb end diff --git a/src/data/fin/basic.lean b/src/data/fin/basic.lean index ec6e542f08e0b..a5f7a3788fc1f 100644 --- a/src/data/fin/basic.lean +++ b/src/data/fin/basic.lean @@ -1192,7 +1192,7 @@ begin end /-- By sending `x` to `last n - x`, `fin n` is order-equivalent to its `order_dual`. -/ -def _root_.order_iso.fin_equiv : ∀ {n}, order_dual (fin n) ≃o fin n +def _root_.order_iso.fin_equiv : ∀ {n}, (fin n)ᵒᵈ ≃o fin n | 0 := ⟨⟨elim0, elim0, elim0, elim0⟩, elim0⟩ | (n+1) := order_iso.symm $ { to_fun := λ x, last n - x, diff --git a/src/data/fin/tuple/basic.lean b/src/data/fin/tuple/basic.lean index c333485f05f3a..97a15f219fdf6 100644 --- a/src/data/fin/tuple/basic.lean +++ b/src/data/fin/tuple/basic.lean @@ -167,7 +167,7 @@ forall_fin_succ.trans $ and_congr iff.rfl $ forall_congr $ λ j, by simp [tail] lemma cons_le [Π i, preorder (α i)] {x : α 0} {q : Π i, α i} {p : Π i : fin n, α i.succ} : cons x p ≤ q ↔ x ≤ q 0 ∧ p ≤ tail q := -@le_cons _ (λ i, order_dual (α i)) _ x q p +@le_cons _ (λ i, (α i)ᵒᵈ) _ x q p lemma cons_le_cons [Π i, preorder (α i)] {x₀ y₀ : α 0} {x y : Π i : fin n, α (i.succ)} : cons x₀ x ≤ cons y₀ y ↔ x₀ ≤ y₀ ∧ x ≤ y := diff --git a/src/data/finset/lattice.lean b/src/data/finset/lattice.lean index 5880a4729d6f1..a594502965903 100644 --- a/src/data/finset/lattice.lean +++ b/src/data/finset/lattice.lean @@ -257,7 +257,7 @@ lemma inf_def : s.inf f = (s.1.map f).inf := rfl fold_empty @[simp] lemma inf_cons {b : β} (h : b ∉ s) : (cons b s h).inf f = f b ⊓ s.inf f := -@sup_cons (order_dual α) _ _ _ _ _ _ h +@sup_cons αᵒᵈ _ _ _ _ _ _ h @[simp] lemma inf_insert [decidable_eq β] {b : β} : (insert b s : finset β).inf f = f b ⊓ s.inf f := fold_insert_idem @@ -274,25 +274,25 @@ fold_map inf_singleton lemma inf_union [decidable_eq β] : (s₁ ∪ s₂).inf f = s₁.inf f ⊓ s₂.inf f := -@sup_union (order_dual α) _ _ _ _ _ _ _ +@sup_union αᵒᵈ _ _ _ _ _ _ _ lemma inf_inf : s.inf (f ⊓ g) = s.inf f ⊓ s.inf g := -@sup_sup (order_dual α) _ _ _ _ _ _ +@sup_sup αᵒᵈ _ _ _ _ _ _ theorem inf_congr {f g : β → α} (hs : s₁ = s₂) (hfg : ∀a∈s₂, f a = g a) : s₁.inf f = s₂.inf g := by subst hs; exact finset.fold_congr hfg @[simp] lemma inf_bUnion [decidable_eq β] (s : finset γ) (t : γ → finset β) : (s.bUnion t).inf f = s.inf (λ x, (t x).inf f) := -@sup_bUnion (order_dual α) _ _ _ _ _ _ _ _ +@sup_bUnion αᵒᵈ _ _ _ _ _ _ _ _ lemma inf_const {s : finset β} (h : s.nonempty) (c : α) : s.inf (λ _, c) = c := -@sup_const (order_dual α) _ _ _ _ h _ +@sup_const αᵒᵈ _ _ _ _ h _ -@[simp] lemma inf_top (s : finset β) : s.inf (λ _, ⊤) = (⊤ : α) := @sup_bot (order_dual α) _ _ _ _ +@[simp] lemma inf_top (s : finset β) : s.inf (λ _, ⊤) = (⊤ : α) := @sup_bot αᵒᵈ _ _ _ _ lemma le_inf_iff {a : α} : a ≤ s.inf f ↔ ∀ b ∈ s, a ≤ f b := -@sup_le_iff (order_dual α) _ _ _ _ _ _ +@sup_le_iff αᵒᵈ _ _ _ _ _ _ lemma inf_le {b : β} (hb : b ∈ s) : s.inf f ≤ f b := le_inf_iff.1 le_rfl _ hb @@ -307,22 +307,22 @@ lemma inf_mono (h : s₁ ⊆ s₂) : s₂.inf f ≤ s₁.inf f := le_inf $ assume b hb, inf_le (h hb) lemma inf_attach (s : finset β) (f : β → α) : s.attach.inf (λ x, f x) = s.inf f := -@sup_attach (order_dual α) _ _ _ _ _ +@sup_attach αᵒᵈ _ _ _ _ _ lemma inf_comm (s : finset β) (t : finset γ) (f : β → γ → α) : s.inf (λ b, t.inf (f b)) = t.inf (λ c, s.inf (λ b, f b c)) := -@sup_comm (order_dual α) _ _ _ _ _ _ _ +@sup_comm αᵒᵈ _ _ _ _ _ _ _ lemma inf_product_left (s : finset β) (t : finset γ) (f : β × γ → α) : (s.product t).inf f = s.inf (λ i, t.inf $ λ i', f ⟨i, i'⟩) := -@sup_product_left (order_dual α) _ _ _ _ _ _ _ +@sup_product_left αᵒᵈ _ _ _ _ _ _ _ lemma inf_product_right (s : finset β) (t : finset γ) (f : β × γ → α) : (s.product t).inf f = t.inf (λ i', s.inf $ λ i, f ⟨i, i'⟩) := -@sup_product_right (order_dual α) _ _ _ _ _ _ _ +@sup_product_right αᵒᵈ _ _ _ _ _ _ _ @[simp] lemma inf_erase_top [decidable_eq α] (s : finset α) : (s.erase ⊤).inf id = s.inf id := -@sup_erase_bot (order_dual α) _ _ _ _ +@sup_erase_bot αᵒᵈ _ _ _ _ lemma sup_sdiff_left {α β : Type*} [boolean_algebra α] (s : finset β) (f : β → α) (a : α) : s.sup (λ b, a \ f b) = a \ s.inf f := @@ -353,18 +353,18 @@ end lemma comp_inf_eq_inf_comp [semilattice_inf γ] [order_top γ] {s : finset β} {f : β → α} (g : α → γ) (g_inf : ∀ x y, g (x ⊓ y) = g x ⊓ g y) (top : g ⊤ = ⊤) : g (s.inf f) = s.inf (g ∘ f) := -@comp_sup_eq_sup_comp (order_dual α) _ (order_dual γ) _ _ _ _ _ _ _ g_inf top +@comp_sup_eq_sup_comp αᵒᵈ _ γᵒᵈ _ _ _ _ _ _ _ g_inf top /-- Computing `inf` in a subtype (closed under `inf`) is the same as computing it in `α`. -/ lemma inf_coe {P : α → Prop} {Ptop : P ⊤} {Pinf : ∀{{x y}}, P x → P y → P (x ⊓ y)} (t : finset β) (f : β → {x : α // P x}) : (@inf _ _ (subtype.semilattice_inf Pinf) (subtype.order_top Ptop) t f : α) = t.inf (λ x, f x) := -@sup_coe (order_dual α) _ _ _ _ Ptop Pinf t f +@sup_coe αᵒᵈ _ _ _ _ Ptop Pinf t f lemma inf_induction {p : α → Prop} (ht : p ⊤) (hp : ∀ a₁, p a₁ → ∀ a₂, p a₂ → p (a₁ ⊓ a₂)) (hs : ∀ b ∈ s, p (f b)) : p (s.inf f) := -@sup_induction (order_dual α) _ _ _ _ _ _ ht hp hs +@sup_induction αᵒᵈ _ _ _ _ _ _ ht hp hs lemma inf_mem (s : set α) (w₁ : ⊤ ∈ s) (w₂ : ∀ x y ∈ s, x ⊓ y ∈ s) @@ -375,7 +375,7 @@ lemma inf_mem @[simp] lemma inf_eq_top_iff (f : β → α) (S : finset β) : S.inf f = ⊤ ↔ ∀ s ∈ S, f s = ⊤ := -@finset.sup_eq_bot_iff (order_dual α) _ _ _ _ _ +@finset.sup_eq_bot_iff αᵒᵈ _ _ _ _ _ end inf @@ -410,11 +410,11 @@ variables [order_top α] lemma inf_sup_distrib_left (s : finset ι) (f : ι → α) (a : α) : a ⊔ s.inf f = s.inf (λ i, a ⊔ f i) := -@sup_inf_distrib_left (order_dual α) _ _ _ _ _ _ +@sup_inf_distrib_left αᵒᵈ _ _ _ _ _ _ lemma inf_sup_distrib_right (s : finset ι) (f : ι → α) (a : α) : s.inf f ⊔ a = s.inf (λ i, f i ⊔ a) := -@sup_inf_distrib_right (order_dual α) _ _ _ _ _ _ +@sup_inf_distrib_right αᵒᵈ _ _ _ _ _ _ end order_top end distrib_lattice @@ -455,22 +455,21 @@ lemma comp_inf_eq_inf_comp_of_is_total [semilattice_inf β] [order_top β] (g : comp_inf_eq_inf_comp g mono_g.map_inf top @[simp] protected lemma inf_le_iff (ha : a < ⊤) : s.inf f ≤ a ↔ ∃ b ∈ s, f b ≤ a := -@finset.le_sup_iff (order_dual α) _ _ _ _ _ _ ha +@finset.le_sup_iff αᵒᵈ _ _ _ _ _ _ ha -@[simp] protected lemma inf_lt_iff : s.inf f < a ↔ (∃ b ∈ s, f b < a) := -@finset.lt_sup_iff (order_dual α) _ _ _ _ _ _ +@[simp] protected lemma inf_lt_iff : s.inf f < a ↔ ∃ b ∈ s, f b < a := +@finset.lt_sup_iff αᵒᵈ _ _ _ _ _ _ @[simp] protected lemma lt_inf_iff (ha : a < ⊤) : a < s.inf f ↔ ∀ b ∈ s, a < f b := -@finset.sup_lt_iff (order_dual α) _ _ _ _ _ _ ha +@finset.sup_lt_iff αᵒᵈ _ _ _ _ _ _ ha end order_top end linear_order -lemma inf_eq_infi [complete_lattice β] (s : finset α) (f : α → β) : s.inf f = (⨅a∈s, f a) := -@sup_eq_supr _ (order_dual β) _ _ _ +lemma inf_eq_infi [complete_lattice β] (s : finset α) (f : α → β) : s.inf f = ⨅ a ∈ s, f a := +@sup_eq_supr _ βᵒᵈ _ _ _ -lemma inf_id_eq_Inf [complete_lattice α] (s : finset α) : s.inf id = Inf s := -@sup_id_eq_Sup (order_dual α) _ _ +lemma inf_id_eq_Inf [complete_lattice α] (s : finset α) : s.inf id = Inf s := @sup_id_eq_Sup αᵒᵈ _ _ lemma inf_id_set_eq_sInter (s : finset (set α)) : s.inf id = ⋂₀(↑s) := inf_id_eq_Inf _ @@ -479,7 +478,7 @@ inf_id_eq_Inf _ inf_eq_infi _ _ lemma inf_eq_Inf_image [complete_lattice β] (s : finset α) (f : α → β) : s.inf f = Inf (f '' s) := -@sup_eq_Sup_image _ (order_dual β) _ _ _ +@sup_eq_Sup_image _ βᵒᵈ _ _ _ section sup' variables [semilattice_sup α] @@ -579,52 +578,48 @@ variables [semilattice_inf α] lemma inf_of_mem {s : finset β} (f : β → α) {b : β} (h : b ∈ s) : ∃ (a : α), s.inf (coe ∘ f : β → with_top α) = ↑a := -@sup_of_mem (order_dual α) _ _ _ f _ h +@sup_of_mem αᵒᵈ _ _ _ f _ h /-- Given nonempty finset `s` then `s.inf' H f` is the infimum of its image under `f` in (possibly unbounded) meet-semilattice `α`, where `H` is a proof of nonemptiness. If `α` has a top element you may instead use `finset.inf` which does not require `s` nonempty. -/ def inf' (s : finset β) (H : s.nonempty) (f : β → α) : α := -@sup' (order_dual α) _ _ s H f +@sup' αᵒᵈ _ _ s H f -variables {s : finset β} (H : s.nonempty) (f : β → α) +variables {s : finset β} (H : s.nonempty) (f : β → α) {a : α} {b : β} @[simp] lemma coe_inf' : ((s.inf' H f : α) : with_top α) = s.inf (coe ∘ f) := -@coe_sup' (order_dual α) _ _ _ H f +@coe_sup' αᵒᵈ _ _ _ H f @[simp] lemma inf'_cons {b : β} {hb : b ∉ s} {h : (cons b s hb).nonempty} : (cons b s hb).inf' h f = f b ⊓ s.inf' H f := -@sup'_cons (order_dual α) _ _ _ H f _ _ _ +@sup'_cons αᵒᵈ _ _ _ H f _ _ _ @[simp] lemma inf'_insert [decidable_eq β] {b : β} {h : (insert b s).nonempty} : (insert b s).inf' h f = f b ⊓ s.inf' H f := -@sup'_insert (order_dual α) _ _ _ H f _ _ _ +@sup'_insert αᵒᵈ _ _ _ H f _ _ _ @[simp] lemma inf'_singleton {b : β} {h : ({b} : finset β).nonempty} : ({b} : finset β).inf' h f = f b := rfl -lemma le_inf' {a : α} (hs : ∀ b ∈ s, a ≤ f b) : a ≤ s.inf' H f := -@sup'_le (order_dual α) _ _ _ H f _ hs - -lemma inf'_le {b : β} (h : b ∈ s) : s.inf' ⟨b, h⟩ f ≤ f b := -@le_sup' (order_dual α) _ _ _ f _ h +lemma le_inf' (hs : ∀ b ∈ s, a ≤ f b) : a ≤ s.inf' H f := @sup'_le αᵒᵈ _ _ _ H f _ hs +lemma inf'_le (h : b ∈ s) : s.inf' ⟨b, h⟩ f ≤ f b := @le_sup' αᵒᵈ _ _ _ f _ h -@[simp] lemma inf'_const (a : α) : s.inf' H (λ b, a) = a := -@sup'_const (order_dual α) _ _ _ _ _ +@[simp] lemma inf'_const (a : α) : s.inf' H (λ b, a) = a := @sup'_const αᵒᵈ _ _ _ _ _ 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) := -@sup'_bUnion (order_dual α) _ _ _ _ _ _ Hs _ Ht +@sup'_bUnion αᵒᵈ _ _ _ _ _ _ Hs _ Ht lemma comp_inf'_eq_inf'_comp [semilattice_inf γ] {s : finset β} (H : s.nonempty) {f : β → α} (g : α → γ) (g_inf : ∀ x y, g (x ⊓ y) = g x ⊓ g y) : g (s.inf' H f) = s.inf' H (g ∘ f) := -@comp_sup'_eq_sup'_comp (order_dual α) _ (order_dual γ) _ _ _ H f g g_inf +@comp_sup'_eq_sup'_comp αᵒᵈ _ γᵒᵈ _ _ _ H f g g_inf lemma inf'_induction {p : α → Prop} (hp : ∀ a₁, p a₁ → ∀ a₂, p a₂ → p (a₁ ⊓ a₂)) (hs : ∀ b ∈ s, p (f b)) : p (s.inf' H f) := -@sup'_induction (order_dual α) _ _ _ H f _ hp hs +@sup'_induction αᵒᵈ _ _ _ H f _ hp hs lemma inf'_mem (s : set α) (w : ∀ x y ∈ s, x ⊓ y ∈ s) {ι : Type*} (t : finset ι) (H : t.nonempty) (p : ι → α) (h : ∀ i ∈ t, p i ∈ s) : @@ -632,7 +627,7 @@ lemma inf'_mem (s : set α) (w : ∀ x y ∈ s, x ⊓ y ∈ s) @[congr] lemma inf'_congr {t : finset β} {f g : β → α} (h₁ : s = t) (h₂ : ∀ x ∈ s, f x = g x) : s.inf' H f = t.inf' (h₁ ▸ H) g := -@sup'_congr (order_dual α) _ _ _ H _ _ _ h₁ h₂ +@sup'_congr αᵒᵈ _ _ _ H _ _ _ h₁ h₂ end inf' @@ -656,15 +651,15 @@ section inf variables [semilattice_inf α] [order_top α] lemma inf'_eq_inf {s : finset β} (H : s.nonempty) (f : β → α) : s.inf' H f = s.inf f := -@sup'_eq_sup (order_dual α) _ _ _ _ H f +@sup'_eq_sup αᵒᵈ _ _ _ _ H f lemma inf_closed_of_inf_closed {s : set α} (t : finset α) (htne : t.nonempty) (h_subset : ↑t ⊆ s) (h : ∀ a b ∈ s, a ⊓ b ∈ s) : t.inf id ∈ s := -@sup_closed_of_sup_closed (order_dual α) _ _ _ t htne h_subset h +@sup_closed_of_sup_closed αᵒᵈ _ _ _ t htne h_subset h lemma coe_inf_of_nonempty {s : finset β} (h : s.nonempty) (f : β → α): (↑(s.inf f) : with_top α) = s.inf (λ i, f i) := -@coe_sup_of_nonempty (order_dual α) _ _ _ _ h f +@coe_sup_of_nonempty αᵒᵈ _ _ _ _ h f end inf @@ -684,7 +679,7 @@ variables {C : β → Type*} [Π (b : β), semilattice_inf (C b)] [Π (b : β), @[simp] protected lemma inf_apply (s : finset α) (f : α → Π (b : β), C b) (b : β) : s.inf f b = s.inf (λ a, f a b) := -@finset.sup_apply _ _ (λ b, order_dual (C b)) _ _ s f b +@finset.sup_apply _ _ (λ b, (C b)ᵒᵈ) _ _ s f b end inf @@ -704,7 +699,7 @@ variables {C : β → Type*} [Π (b : β), semilattice_inf (C b)] @[simp] protected lemma inf'_apply {s : finset α} (H : s.nonempty) (f : α → Π (b : β), C b) (b : β) : s.inf' H f b = s.inf' H (λ a, f a b) := -@finset.sup'_apply _ _ (λ b, order_dual (C b)) _ _ H f b +@finset.sup'_apply _ _ (λ b, (C b)ᵒᵈ) _ _ H f b end inf' @@ -729,14 +724,9 @@ begin exact ball_congr (λ b hb, with_bot.coe_lt_coe), end -@[simp] lemma inf'_le_iff : s.inf' H f ≤ a ↔ ∃ i ∈ s, f i ≤ a := -@le_sup'_iff (order_dual α) _ _ _ H f _ - -@[simp] lemma inf'_lt_iff : s.inf' H f < a ↔ ∃ i ∈ s, f i < a := -@lt_sup'_iff (order_dual α) _ _ _ H f _ - -@[simp] lemma lt_inf'_iff : a < s.inf' H f ↔ ∀ i ∈ s, a < f i := -@sup'_lt_iff (order_dual α) _ _ _ H f _ +@[simp] lemma inf'_le_iff : s.inf' H f ≤ a ↔ ∃ i ∈ s, f i ≤ a := @le_sup'_iff αᵒᵈ _ _ _ H f _ +@[simp] lemma inf'_lt_iff : s.inf' H f < a ↔ ∃ i ∈ s, f i < a := @lt_sup'_iff αᵒᵈ _ _ _ H f _ +@[simp] lemma lt_inf'_iff : a < s.inf' H f ↔ ∀ i ∈ s, a < f i := @sup'_lt_iff αᵒᵈ _ _ _ H f _ lemma exists_mem_eq_sup' (f : ι → α) : ∃ i, i ∈ s ∧ s.sup' H f = f i := begin @@ -750,7 +740,7 @@ begin end lemma exists_mem_eq_inf' (f : ι → α) : ∃ i, i ∈ s ∧ s.inf' H f = f i := -@exists_mem_eq_sup' (order_dual α) _ _ _ H f +@exists_mem_eq_sup' αᵒᵈ _ _ _ H f lemma exists_mem_eq_sup [order_bot α] (s : finset ι) (h : s.nonempty) (f : ι → α) : ∃ i, i ∈ s ∧ s.sup f = f i := @@ -758,7 +748,7 @@ sup'_eq_sup h f ▸ exists_mem_eq_sup' h f lemma exists_mem_eq_inf [order_top α] (s : finset ι) (h : s.nonempty) (f : ι → α) : ∃ i, i ∈ s ∧ s.inf f = f i := -@exists_mem_eq_sup (order_dual α) _ _ _ _ h f +@exists_mem_eq_sup αᵒᵈ _ _ _ _ h f end linear_order @@ -839,8 +829,7 @@ theorem min_eq_none {s : finset α} : s.min = none ↔ s = ∅ := (λ H, let ⟨a, ha⟩ := min_of_nonempty H in by rw h at ha; cases ha), λ h, h.symm ▸ min_empty⟩ -theorem mem_of_min {s : finset α} : ∀ {a : α}, a ∈ s.min → a ∈ s := -@mem_of_max (order_dual α) _ s +theorem mem_of_min {s : finset α} : ∀ {a : α}, a ∈ s.min → a ∈ s := @mem_of_max αᵒᵈ _ s theorem min_le_of_mem {s : finset α} {a b : α} (h₁ : b ∈ s) (h₂ : a ∈ s.min) : a ≤ b := by rcases @inf_le (with_top α) _ _ _ _ _ _ h₁ _ rfl with ⟨b', hb, ab⟩; @@ -862,7 +851,7 @@ def max' (s : finset α) (H : s.nonempty) : α := let ⟨k, hk⟩ := H in let ⟨b, hb⟩ := max_of_mem hk in by simp at hb; simp [hb] -variables (s : finset α) (H : s.nonempty) +variables (s : finset α) (H : s.nonempty) {x : α} theorem min'_mem : s.min' H ∈ s := mem_of_min $ by simp [min'] @@ -894,14 +883,12 @@ is_lub_le_iff (is_greatest_max' s H).is_lub @[simp] lemma max'_lt_iff {x} : s.max' H < x ↔ ∀ y ∈ s, y < x := ⟨λ Hlt y hy, (s.le_max' y hy).trans_lt Hlt, λ H, H _ $ s.max'_mem _⟩ -@[simp] lemma lt_min'_iff {x} : x < s.min' H ↔ ∀ y ∈ s, x < y := -@max'_lt_iff (order_dual α) _ _ H _ +@[simp] lemma lt_min'_iff : x < s.min' H ↔ ∀ y ∈ s, x < y := @max'_lt_iff αᵒᵈ _ _ H _ lemma max'_eq_sup' : s.max' H = s.sup' H id := eq_of_forall_ge_iff $ λ a, (max'_le_iff _ _).trans (sup'_le_iff _ _).symm -lemma min'_eq_inf' : s.min' H = s.inf' H id := -@max'_eq_sup' (order_dual α) _ s H +lemma min'_eq_inf' : s.min' H = s.inf' H id := @max'_eq_sup' αᵒᵈ _ s H /-- `{a}.max' _` is `a`. -/ @[simp] lemma max'_singleton (a : α) : @@ -970,7 +957,7 @@ lt_of_le_of_ne (le_max' _ _ (mem_of_mem_erase ha)) $ ne_of_mem_of_not_mem ha $ n lemma min'_lt_of_mem_erase_min' [decidable_eq α] {a : α} (ha : a ∈ s.erase (s.min' H)) : s.min' H < a := -@lt_max'_of_mem_erase_max' (order_dual α) _ s H _ a ha +@lt_max'_of_mem_erase_max' αᵒᵈ _ s H _ a ha @[simp] lemma max'_image [linear_order β] {f : α → β} (hf : monotone f) (s : finset α) (h : (s.image f).nonempty) : @@ -1019,7 +1006,7 @@ end @[elab_as_eliminator] lemma induction_on_min [decidable_eq α] {p : finset α → Prop} (s : finset α) (h0 : p ∅) (step : ∀ a s, (∀ x ∈ s, a < x) → p s → p (insert a s)) : p s := -@induction_on_max (order_dual α) _ _ _ s h0 step +@induction_on_max αᵒᵈ _ _ _ s h0 step end max_min @@ -1061,7 +1048,7 @@ ordered type : a predicate is true on all `s : finset α` provided that: lemma induction_on_min_value [decidable_eq ι] (f : ι → α) {p : finset ι → Prop} (s : finset ι) (h0 : p ∅) (step : ∀ a s, a ∉ s → (∀ x ∈ s, f a ≤ f x) → p s → p (insert a s)) : p s := -@induction_on_max_value (order_dual α) ι _ _ _ _ s h0 step +@induction_on_max_value αᵒᵈ ι _ _ _ _ s h0 step end max_min_induction_value @@ -1078,7 +1065,7 @@ end lemma exists_min_image (s : finset β) (f : β → α) (h : s.nonempty) : ∃ x ∈ s, ∀ x' ∈ s, f x ≤ f x' := -@exists_max_image (order_dual α) β _ s f h +@exists_max_image αᵒᵈ β _ s f h end exists_max_min end finset @@ -1175,16 +1162,15 @@ by rw [← supr_eq_supr_finset, ← equiv.plift.surjective.supr_comp]; refl /-- Infimum of `s i`, `i : ι`, is equal to the infimum over `t : finset ι` of infima `⨅ i ∈ t, s i`. This version assumes `ι` is a `Type*`. See `infi_eq_infi_finset'` for a version that works for `ι : Sort*`. -/ -lemma infi_eq_infi_finset (s : ι → α) : - (⨅i, s i) = (⨅t:finset ι, ⨅i∈t, s i) := -@supr_eq_supr_finset (order_dual α) _ _ _ +lemma infi_eq_infi_finset (s : ι → α) : (⨅ i, s i) = ⨅ (t : finset ι) (i ∈ t), s i := +@supr_eq_supr_finset αᵒᵈ _ _ _ /-- Infimum of `s i`, `i : ι`, is equal to the infimum over `t : finset ι` of infima `⨅ i ∈ t, s i`. This version works for `ι : Sort*`. See `infi_eq_infi_finset` for a version that assumes `ι : Type*` but has no `plift`s. -/ lemma infi_eq_infi_finset' (s : ι' → α) : (⨅i, s i) = (⨅t:finset (plift ι'), ⨅i∈t, s (plift.down i)) := -@supr_eq_supr_finset' (order_dual α) _ _ _ +@supr_eq_supr_finset' αᵒᵈ _ _ _ end lattice @@ -1273,9 +1259,8 @@ lemma supr_option_to_finset (o : option α) (f : α → β) : (⨆ x ∈ o.to_finset, f x) = ⨆ x ∈ o, f x := by simp -lemma infi_option_to_finset (o : option α) (f : α → β) : - (⨅ x ∈ o.to_finset, f x) = ⨅ x ∈ o, f x := -@supr_option_to_finset _ (order_dual β) _ _ _ +lemma infi_option_to_finset (o : option α) (f : α → β) : (⨅ x ∈ o.to_finset, f x) = ⨅ x ∈ o, f x := +@supr_option_to_finset _ βᵒᵈ _ _ _ variables [decidable_eq α] @@ -1285,7 +1270,7 @@ by simp [supr_or, supr_sup_eq] theorem infi_union {f : α → β} {s t : finset α} : (⨅ x ∈ s ∪ t, f x) = (⨅ x ∈ s, f x) ⊓ (⨅ x ∈ t, f x) := -@supr_union α (order_dual β) _ _ _ _ _ +@supr_union α βᵒᵈ _ _ _ _ _ lemma supr_insert (a : α) (s : finset α) (t : α → β) : (⨆ x ∈ insert a s, t x) = t a ⊔ (⨆ x ∈ s, t x) := @@ -1293,7 +1278,7 @@ by { rw insert_eq, simp only [supr_union, finset.supr_singleton] } lemma infi_insert (a : α) (s : finset α) (t : α → β) : (⨅ x ∈ insert a s, t x) = t a ⊓ (⨅ x ∈ s, t x) := -@supr_insert α (order_dual β) _ _ _ _ _ +@supr_insert α βᵒᵈ _ _ _ _ _ lemma supr_finset_image {f : γ → α} {g : α → β} {s : finset γ} : (⨆ x ∈ s.image f, g x) = (⨆ y ∈ s, g (f y)) := @@ -1320,7 +1305,7 @@ end lemma infi_insert_update {x : α} {t : finset α} (f : α → β) {s : β} (hx : x ∉ t) : (⨅ (i ∈ insert x t), update f x s i) = (s ⊓ ⨅ (i ∈ t), f i) := -@supr_insert_update α (order_dual β) _ _ _ _ f _ hx +@supr_insert_update α βᵒᵈ _ _ _ _ f _ hx lemma supr_bUnion (s : finset γ) (t : γ → finset α) (f : α → β) : (⨆ y ∈ s.bUnion t, f y) = ⨆ (x ∈ s) (y ∈ t x), f y := @@ -1328,7 +1313,7 @@ by simp [@supr_comm _ α, supr_and] lemma infi_bUnion (s : finset γ) (t : γ → finset α) (f : α → β) : (⨅ y ∈ s.bUnion t, f y) = ⨅ (x ∈ s) (y ∈ t x), f y := -@supr_bUnion _ (order_dual β) _ _ _ _ _ _ +@supr_bUnion _ βᵒᵈ _ _ _ _ _ _ end lattice diff --git a/src/data/finset/locally_finite.lean b/src/data/finset/locally_finite.lean index 462782b038ba2..ebe1b669ddf15 100644 --- a/src/data/finset/locally_finite.lean +++ b/src/data/finset/locally_finite.lean @@ -316,7 +316,7 @@ begin end lemma card_Ioc_eq_card_Icc_sub_one (a b : α) : (Ioc a b).card = (Icc a b).card - 1 := -@card_Ico_eq_card_Icc_sub_one (order_dual α) _ _ _ _ +@card_Ico_eq_card_Icc_sub_one αᵒᵈ _ _ _ _ lemma card_Ioo_eq_card_Ico_sub_one (a b : α) : (Ioo a b).card = (Ico a b).card - 1 := begin @@ -330,7 +330,7 @@ begin end lemma card_Ioo_eq_card_Ioc_sub_one (a b : α) : (Ioo a b).card = (Ioc a b).card - 1 := -@card_Ioo_eq_card_Ico_sub_one (order_dual α) _ _ _ _ +@card_Ioo_eq_card_Ico_sub_one αᵒᵈ _ _ _ _ lemma card_Ioo_eq_card_Icc_sub_two (a b : α) : (Ioo a b).card = (Icc a b).card - 2 := by { rw [card_Ioo_eq_card_Ico_sub_one, card_Ico_eq_card_Icc_sub_one], refl } diff --git a/src/data/finset/pi_induction.lean b/src/data/finset/pi_induction.lean index 8cf3ac1a2f5b1..23fd8ef8cb3bb 100644 --- a/src/data/finset/pi_induction.lean +++ b/src/data/finset/pi_induction.lean @@ -92,6 +92,6 @@ lemma induction_on_pi_min [Π i, linear_order (α i)] {p : (Π i, finset (α i)) (step : ∀ (g : Π i, finset (α i)) (i : ι) (x : α i), (∀ y ∈ g i, x < y) → p g → p (update g i (insert x (g i)))) : p f := -@induction_on_pi_max ι (λ i, order_dual (α i)) _ _ _ _ _ _ h0 step +@induction_on_pi_max ι (λ i, (α i)ᵒᵈ) _ _ _ _ _ _ h0 step end finset diff --git a/src/data/finset/sigma.lean b/src/data/finset/sigma.lean index 9456e4dfb452d..6b24cf5fc1a91 100644 --- a/src/data/finset/sigma.lean +++ b/src/data/finset/sigma.lean @@ -71,7 +71,7 @@ end lemma inf_sigma [semilattice_inf β] [order_top β] : (s.sigma t).inf f = s.inf (λ i, (t i).inf $ λ b, f ⟨i, b⟩) := -@sup_sigma _ _ (order_dual β) _ _ _ _ _ +@sup_sigma _ _ βᵒᵈ _ _ _ _ _ end sigma diff --git a/src/data/fintype/basic.lean b/src/data/fintype/basic.lean index 6a0ad4418444b..f323070c5ba65 100644 --- a/src/data/fintype/basic.lean +++ b/src/data/fintype/basic.lean @@ -218,7 +218,7 @@ lemma sup_univ_eq_supr [complete_lattice β] (f : α → β) : finset.univ.sup f /-- A special case of `finset.inf_eq_infi` that omits the useless `x ∈ univ` binder. -/ lemma inf_univ_eq_infi [complete_lattice β] (f : α → β) : finset.univ.inf f = infi f := -sup_univ_eq_supr (by exact f : α → order_dual β) +sup_univ_eq_supr (by exact f : α → βᵒᵈ) @[simp] lemma fold_inf_univ [semilattice_inf α] [order_bot α] (a : α) : finset.univ.fold (⊓) a (λ x, x) = ⊥ := @@ -227,7 +227,7 @@ eq_bot_iff.2 $ ((finset.fold_op_rel_iff_and $ @_root_.le_inf_iff α _).1 le_rfl) @[simp] lemma fold_sup_univ [semilattice_sup α] [order_top α] (a : α) : finset.univ.fold (⊔) a (λ x, x) = ⊤ := -@fold_inf_univ (order_dual α) ‹fintype α› _ _ _ +@fold_inf_univ αᵒᵈ ‹fintype α› _ _ _ end finset @@ -891,10 +891,10 @@ fintype.of_equiv _ equiv.plift.symm fintype.card (plift α) = fintype.card α := fintype.of_equiv_card _ -instance (α : Type*) [fintype α] : fintype (order_dual α) := ‹fintype α› +instance (α : Type*) [fintype α] : fintype αᵒᵈ := ‹fintype α› -@[simp] lemma fintype.card_order_dual (α : Type*) [fintype α] : - fintype.card (order_dual α) = fintype.card α := rfl +@[simp] lemma fintype.card_order_dual (α : Type*) [fintype α] : fintype.card αᵒᵈ = fintype.card α := +rfl instance (α : Type*) [fintype α] : fintype (lex α) := ‹fintype α› @@ -1788,7 +1788,7 @@ end lemma finset.exists_maximal {α : Type*} [preorder α] (s : finset α) (h : s.nonempty) : ∃ m ∈ s, ∀ x ∈ s, ¬ (m < x) := -@finset.exists_minimal (order_dual α) _ s h +@finset.exists_minimal αᵒᵈ _ s h namespace infinite diff --git a/src/data/list/big_operators.lean b/src/data/list/big_operators.lean index f1b6b382975a9..bb70195dae307 100644 --- a/src/data/list/big_operators.lean +++ b/src/data/list/big_operators.lean @@ -235,7 +235,7 @@ lemma pow_card_le_prod [preorder M] [covariant_class M M (function.swap (*)) (≤)] [covariant_class M M (*) (≤)] (l : list M) (n : M) (h : ∀ (x ∈ l), n ≤ x) : n ^ l.length ≤ l.prod := -@prod_le_pow_card (order_dual M) _ _ _ _ l n h +@prod_le_pow_card Mᵒᵈ _ _ _ _ l n h @[to_additive exists_lt_of_sum_lt] lemma exists_lt_of_prod_lt' [linear_order M] [covariant_class M M (function.swap (*)) (≤)] [covariant_class M M (*) (≤)] {l : list ι} diff --git a/src/data/list/min_max.lean b/src/data/list/min_max.lean index f38af432cc5ec..c79f23e5b76c8 100644 --- a/src/data/list/min_max.lean +++ b/src/data/list/min_max.lean @@ -34,8 +34,7 @@ l.foldl (argmax₂ f) none /-- `argmin f l` returns `some a`, where `a` of `l` that minimises `f a`. If there are `a b` such that `f a = f b`, it returns whichever of `a` or `b` comes first in the list. `argmin f []` = none` -/ -def argmin (f : α → β) (l : list α) := -@argmax _ (order_dual β) _ f l +def argmin (f : α → β) (l : list α) := @argmax _ βᵒᵈ _ f l @[simp] lemma argmax_two_self (f : α → β) (a : α) : argmax₂ f (some a) a = a := if_pos le_rfl @@ -98,19 +97,19 @@ theorem argmax_mem {f : α → β} : Π {l : list α} {m : α}, m ∈ argmax f l | (hd::tl) m := by simpa [argmax, argmax₂] using foldl_argmax₂_mem f tl hd m theorem argmin_mem {f : α → β} : Π {l : list α} {m : α}, m ∈ argmin f l → m ∈ l := -@argmax_mem _ (order_dual β) _ _ +@argmax_mem _ βᵒᵈ _ _ @[simp] theorem argmax_eq_none {f : α → β} {l : list α} : l.argmax f = none ↔ l = [] := by simp [argmax] @[simp] theorem argmin_eq_none {f : α → β} {l : list α} : l.argmin f = none ↔ l = [] := -@argmax_eq_none _ (order_dual β) _ _ _ +@argmax_eq_none _ βᵒᵈ _ _ _ theorem le_argmax_of_mem {f : α → β} {a m : α} {l : list α} : a ∈ l → m ∈ argmax f l → f a ≤ f m := le_of_foldl_argmax₂ theorem argmin_le_of_mem {f : α → β} {a m : α} {l : list α} : a ∈ l → m ∈ argmin f l → f m ≤ f a:= -@le_argmax_of_mem _ (order_dual β) _ _ _ _ _ +@le_argmax_of_mem _ βᵒᵈ _ _ _ _ _ theorem argmax_concat (f : α → β) (a : α) (l : list α) : argmax f (l ++ [a]) = option.cases_on (argmax f l) (some a) (λ c, if f a ≤ f c then some c else some a) := @@ -118,7 +117,7 @@ by rw [argmax, argmax]; simp [argmax₂] theorem argmin_concat (f : α → β) (a : α) (l : list α) : argmin f (l ++ [a]) = option.cases_on (argmin f l) (some a) (λ c, if f c ≤ f a then some c else some a) := -@argmax_concat _ (order_dual β) _ _ _ _ +@argmax_concat _ βᵒᵈ _ _ _ _ theorem argmax_cons (f : α → β) (a : α) (l : list α) : argmax f (a :: l) = option.cases_on (argmax f l) (some a) (λ c, if f c ≤ f a then some a else some c) := @@ -144,7 +143,7 @@ list.reverse_rec_on l rfl $ theorem argmin_cons (f : α → β) (a : α) (l : list α) : argmin f (a :: l) = option.cases_on (argmin f l) (some a) (λ c, if f a ≤ f c then some a else some c) := -@argmax_cons _ (order_dual β) _ _ _ _ +@argmax_cons _ βᵒᵈ _ _ _ _ theorem index_of_argmax [decidable_eq α] {f : α → β} : Π {l : list α} {m : α}, m ∈ argmax f l → ∀ {a}, a ∈ l → f m ≤ f a → l.index_of m ≤ l.index_of a @@ -169,7 +168,7 @@ end theorem index_of_argmin [decidable_eq α] {f : α → β} : Π {l : list α} {m : α}, m ∈ argmin f l → ∀ {a}, a ∈ l → f a ≤ f m → l.index_of m ≤ l.index_of a := -@index_of_argmax _ (order_dual β) _ _ _ +@index_of_argmax _ βᵒᵈ _ _ _ theorem mem_argmax_iff [decidable_eq α] {f : α → β} {m : α} {l : list α} : m ∈ argmax f l ↔ m ∈ l ∧ (∀ a ∈ l, f a ≤ f m) ∧ @@ -191,7 +190,7 @@ theorem argmax_eq_some_iff [decidable_eq α] {f : α → β} {m : α} {l : list theorem mem_argmin_iff [decidable_eq α] {f : α → β} {m : α} {l : list α} : m ∈ argmin f l ↔ m ∈ l ∧ (∀ a ∈ l, f m ≤ f a) ∧ (∀ a ∈ l, f a ≤ f m → l.index_of m ≤ l.index_of a) := -@mem_argmax_iff _ (order_dual β) _ _ _ _ _ +@mem_argmax_iff _ βᵒᵈ _ _ _ _ _ theorem argmin_eq_some_iff [decidable_eq α] {f : α → β} {m : α} {l : list α} : argmin f l = some m ↔ m ∈ l ∧ (∀ a ∈ l, f m ≤ f a) ∧ @@ -236,7 +235,7 @@ option.cases_on (maximum l) (λ _ h, absurd ha ((h rfl).symm ▸ not_mem_nil _)) (@maximum_eq_none _ _ l).1 theorem le_minimum_of_mem' {a : α} {l : list α} (ha : a ∈ l) : minimum l ≤ (a : with_top α) := -@le_maximum_of_mem' (order_dual α) _ _ _ ha +@le_maximum_of_mem' αᵒᵈ _ _ _ ha theorem maximum_concat (a : α) (l : list α) : maximum (l ++ [a]) = max (maximum l) a := begin @@ -250,14 +249,14 @@ begin end theorem minimum_concat (a : α) (l : list α) : minimum (l ++ [a]) = min (minimum l) a := -@maximum_concat (order_dual α) _ _ _ +@maximum_concat αᵒᵈ _ _ _ theorem maximum_cons (a : α) (l : list α) : maximum (a :: l) = max a (maximum l) := list.reverse_rec_on l (by simp [@max_eq_left (with_bot α) _ _ _ bot_le]) (λ tl hd ih, by rw [← cons_append, maximum_concat, ih, maximum_concat, max_assoc]) theorem minimum_cons (a : α) (l : list α) : minimum (a :: l) = min a (minimum l) := -@maximum_cons (order_dual α) _ _ _ +@maximum_cons αᵒᵈ _ _ _ theorem maximum_eq_coe_iff {m : α} {l : list α} : maximum l = m ↔ m ∈ l ∧ (∀ a ∈ l, a ≤ m) := @@ -273,7 +272,7 @@ end theorem minimum_eq_coe_iff {m : α} {l : list α} : minimum l = m ↔ m ∈ l ∧ (∀ a ∈ l, m ≤ a) := -@maximum_eq_coe_iff (order_dual α) _ _ _ +@maximum_eq_coe_iff αᵒᵈ _ _ _ section fold @@ -281,7 +280,7 @@ variables {M : Type*} [canonically_linear_ordered_add_monoid M] /-! Note: since there is no typeclass typeclass dual to `canonically_linear_ordered_add_monoid α` we cannot express these lemmas generally for -`minimum`; instead we are limited to doing so on `order_dual α`. -/ +`minimum`; instead we are limited to doing so on `αᵒᵈ`. -/ lemma maximum_eq_coe_foldr_max_of_ne_nil (l : list M) (h : l ≠ []) : l.maximum = (l.foldr max ⊥ : M) := @@ -294,8 +293,8 @@ begin { simp [IH h] } } end -lemma minimum_eq_coe_foldr_min_of_ne_nil (l : list (order_dual M)) (h : l ≠ []) : - l.minimum = (l.foldr min ⊤ : order_dual M) := +lemma minimum_eq_coe_foldr_min_of_ne_nil (l : list Mᵒᵈ) (h : l ≠ []) : + l.minimum = (l.foldr min ⊤ : Mᵒᵈ) := maximum_eq_coe_foldr_max_of_ne_nil l h lemma maximum_nat_eq_coe_foldr_max_of_ne_nil (l : list ℕ) (h : l ≠ []) : @@ -312,9 +311,7 @@ begin simpa [hy] using IH } end -lemma le_min_of_le_forall (l : list (order_dual M)) (n : (order_dual M)) - (h : ∀ (x ∈ l), n ≤ x) : - n ≤ l.foldr min ⊤ := +lemma le_min_of_le_forall (l : list Mᵒᵈ) (n : Mᵒᵈ) (h : ∀ (x ∈ l), n ≤ x) : n ≤ l.foldr min ⊤ := max_le_of_forall_le l n h lemma max_nat_le_of_forall_le (l : list ℕ) (n : ℕ) (h : ∀ (x ∈ l), x ≤ n) : diff --git a/src/data/nat/lattice.lean b/src/data/nat/lattice.lean index f6341a98e5162..2e0578ba30d21 100644 --- a/src/data/nat/lattice.lean +++ b/src/data/nat/lattice.lean @@ -152,10 +152,10 @@ lemma supr_lt_succ' (u : ℕ → α) (n : ℕ) : (⨆ k < n + 1, u k) = u 0 ⊔ by { rw ← sup_supr_nat_succ, simp } lemma infi_lt_succ (u : ℕ → α) (n : ℕ) : (⨅ k < n + 1, u k) = (⨅ k < n, u k) ⊓ u n := -@supr_lt_succ (order_dual α) _ _ _ +@supr_lt_succ αᵒᵈ _ _ _ lemma infi_lt_succ' (u : ℕ → α) (n : ℕ) : (⨅ k < n + 1, u k) = u 0 ⊓ (⨅ k < n, u (k + 1)) := -@supr_lt_succ' (order_dual α) _ _ _ +@supr_lt_succ' αᵒᵈ _ _ _ end diff --git a/src/data/nat/pairing.lean b/src/data/nat/pairing.lean index 365d0baaa34f4..3b61b8826f839 100644 --- a/src/data/nat/pairing.lean +++ b/src/data/nat/pairing.lean @@ -140,7 +140,7 @@ by rw [← (supr_prod : (⨆ i : ℕ × ℕ, f i.1 i.2) = _), ← nat.surjective lemma infi_unpair {α} [complete_lattice α] (f : ℕ → ℕ → α) : (⨅ n : ℕ, f n.unpair.1 n.unpair.2) = ⨅ i j : ℕ, f i j := -supr_unpair (show ℕ → ℕ → order_dual α, from f) +supr_unpair (show ℕ → ℕ → αᵒᵈ, from f) end complete_lattice diff --git a/src/data/ordmap/ordset.lean b/src/data/ordmap/ordset.lean index a19fea78f2afa..b0aacf4dd1846 100644 --- a/src/data/ordmap/ordset.lean +++ b/src/data/ordmap/ordset.lean @@ -540,7 +540,7 @@ theorem find_max'_all {P : α → Prop} : ∀ (x : α) t, P x → all P t → P /-! ### `insert` -/ theorem dual_insert [preorder α] [is_total α (≤)] [@decidable_rel α (≤)] (x : α) : - ∀ t : ordnode α, dual (ordnode.insert x t) = @ordnode.insert (order_dual α) _ _ x (dual t) + ∀ t : ordnode α, dual (ordnode.insert x t) = @ordnode.insert αᵒᵈ _ _ x (dual t) | nil := rfl | (node _ l y r) := begin rw [ordnode.insert, dual, ordnode.insert, order_dual.cmp_le_flip, ← cmp_le_swap x y], @@ -796,12 +796,12 @@ def bounded : ordnode α → with_bot α → with_top α → Prop | (node _ l x r) o₁ o₂ := bounded l o₁ ↑x ∧ bounded r ↑x o₂ theorem bounded.dual : ∀ {t : ordnode α} {o₁ o₂} (h : bounded t o₁ o₂), - @bounded (order_dual α) _ (dual t) o₂ o₁ + @bounded αᵒᵈ _ (dual t) o₂ o₁ | nil o₁ o₂ h := by cases o₁; cases o₂; try {trivial}; exact h | (node s l x r) _ _ ⟨ol, or⟩ := ⟨or.dual, ol.dual⟩ -theorem bounded.dual_iff {t : ordnode α} {o₁ o₂} : bounded t o₁ o₂ ↔ - @bounded (order_dual α) _ (dual t) o₂ o₁ := +theorem bounded.dual_iff {t : ordnode α} {o₁ o₂} : + bounded t o₁ o₂ ↔ @bounded αᵒᵈ _ (dual t) o₂ o₁ := ⟨bounded.dual, λ h, by have := bounded.dual h; rwa [dual_dual, order_dual.preorder.dual_dual] at this⟩ @@ -927,8 +927,7 @@ theorem valid'.node {s l x r o₁ o₂} valid' o₁ (@node α s l x r) o₂ := ⟨⟨hl.1, hr.1⟩, ⟨hs, hl.2, hr.2⟩, ⟨H, hl.3, hr.3⟩⟩ -theorem valid'.dual : ∀ {t : ordnode α} {o₁ o₂} (h : valid' o₁ t o₂), - @valid' (order_dual α) _ o₂ (dual t) o₁ +theorem valid'.dual : ∀ {t : ordnode α} {o₁ o₂} (h : valid' o₁ t o₂), @valid' αᵒᵈ _ o₂ (dual t) o₁ | nil o₁ o₂ h := valid'_nil h.1.dual | (node s l x r) o₁ o₂ ⟨⟨ol, or⟩, ⟨rfl, sl, sr⟩, ⟨b, bl, br⟩⟩ := let ⟨ol', sl', bl'⟩ := valid'.dual ⟨ol, sl, bl⟩, @@ -937,16 +936,13 @@ theorem valid'.dual : ∀ {t : ordnode α} {o₁ o₂} (h : valid' o₁ t o₂), ⟨by simp [size_dual, add_comm], sr', sl'⟩, ⟨by rw [size_dual, size_dual]; exact b.symm, br', bl'⟩⟩ -theorem valid'.dual_iff {t : ordnode α} {o₁ o₂} : valid' o₁ t o₂ ↔ - @valid' (order_dual α) _ o₂ (dual t) o₁ := +theorem valid'.dual_iff {t : ordnode α} {o₁ o₂} : valid' o₁ t o₂ ↔ @valid' αᵒᵈ _ o₂ (dual t) o₁ := ⟨valid'.dual, λ h, by have := valid'.dual h; rwa [dual_dual, order_dual.preorder.dual_dual] at this⟩ -theorem valid.dual {t : ordnode α} : valid t → - @valid (order_dual α) _ (dual t) := valid'.dual +theorem valid.dual {t : ordnode α} : valid t → @valid αᵒᵈ _ (dual t) := valid'.dual -theorem valid.dual_iff {t : ordnode α} : valid t ↔ - @valid (order_dual α) _ (dual t) := valid'.dual_iff +theorem valid.dual_iff {t : ordnode α} : valid t ↔ @valid αᵒᵈ _ (dual t) := valid'.dual_iff theorem valid'.left {s l x r o₁ o₂} (H : valid' o₁ (@node α s l x r) o₂) : valid' o₁ l x := ⟨H.1.1, H.2.2.1, H.3.2.1⟩ diff --git a/src/data/real/ennreal.lean b/src/data/real/ennreal.lean index 001e3031e1c04..db0e3a5ab3140 100644 --- a/src/data/real/ennreal.lean +++ b/src/data/real/ennreal.lean @@ -253,10 +253,10 @@ lemma infi_ne_top [complete_lattice α] (f : ℝ≥0∞ → α) : (⨅ x ≠ ∞ by rw [infi_subtype', cinfi_ne_top] lemma csupr_ne_top [has_Sup α] (f : ℝ≥0∞ → α) : (⨆ x : {x // x ≠ ∞}, f x) = ⨆ x : ℝ≥0, f x := -@cinfi_ne_top (order_dual α) _ _ +@cinfi_ne_top αᵒᵈ _ _ lemma supr_ne_top [complete_lattice α] (f : ℝ≥0∞ → α) : (⨆ x ≠ ∞, f x) = ⨆ x : ℝ≥0, f x := -@infi_ne_top (order_dual α) _ _ +@infi_ne_top αᵒᵈ _ _ lemma infi_ennreal {α : Type*} [complete_lattice α] {f : ℝ≥0∞ → α} : (⨅ n, f n) = (⨅ n : ℝ≥0, f n) ⊓ f ∞ := @@ -266,7 +266,7 @@ le_antisymm lemma supr_ennreal {α : Type*} [complete_lattice α] {f : ℝ≥0∞ → α} : (⨆ n, f n) = (⨆ n : ℝ≥0, f n) ⊔ f ∞ := -@infi_ennreal (order_dual α) _ _ +@infi_ennreal αᵒᵈ _ _ @[simp] lemma add_top : a + ∞ = ∞ := add_top _ @[simp] lemma top_add : ∞ + a = ∞ := top_add _ @@ -1068,7 +1068,7 @@ inv_lt_iff_inv_lt.trans $ by rw [inv_one] /-- The inverse map `λ x, x⁻¹` is an order isomorphism between `ℝ≥0∞` and its `order_dual` -/ @[simps apply] -def _root_.order_iso.inv_ennreal : ℝ≥0∞ ≃o order_dual ℝ≥0∞ := +def _root_.order_iso.inv_ennreal : ℝ≥0∞ ≃o ℝ≥0∞ᵒᵈ := { to_fun := λ x, x⁻¹, inv_fun := λ x, x⁻¹, map_rel_iff' := λ a b, ennreal.inv_le_inv, diff --git a/src/data/real/ereal.lean b/src/data/real/ereal.lean index aec1b4bd4bbcc..c1c818caae5b0 100644 --- a/src/data/real/ereal.lean +++ b/src/data/real/ereal.lean @@ -415,7 +415,7 @@ by conv_lhs { rw [ereal.neg_le, neg_neg] } @[simp, norm_cast] lemma coe_neg (x : ℝ) : ((- x : ℝ) : ereal) = - (x : ereal) := rfl /-- Negation as an order reversing isomorphism on `ereal`. -/ -def neg_order_iso : ereal ≃o (order_dual ereal) := +def neg_order_iso : ereal ≃o erealᵒᵈ := { to_fun := λ x, order_dual.to_dual (-x), inv_fun := λ x, -x.of_dual, map_rel_iff' := λ x y, neg_le_neg_iff, diff --git a/src/data/set/finite.lean b/src/data/set/finite.lean index 01f7b87e09b23..de6dfea7e0e91 100644 --- a/src/data/set/finite.lean +++ b/src/data/set/finite.lean @@ -845,13 +845,12 @@ section variables [semilattice_inf α] [nonempty α] {s : set α} /--A finite set is bounded below.-/ -protected lemma finite.bdd_below (hs : finite s) : bdd_below s := -@finite.bdd_above (order_dual α) _ _ _ hs +protected lemma finite.bdd_below (hs : finite s) : bdd_below s := @finite.bdd_above αᵒᵈ _ _ _ hs /--A finite union of sets which are all bounded below is still bounded below.-/ lemma finite.bdd_below_bUnion {I : set β} {S : β → set α} (H : finite I) : - (bdd_below (⋃i∈I, S i)) ↔ (∀i ∈ I, bdd_below (S i)) := -@finite.bdd_above_bUnion (order_dual α) _ _ _ _ _ H + bdd_below (⋃ i ∈ I, S i) ↔ ∀ i ∈ I, bdd_below (S i) := +@finite.bdd_above_bUnion αᵒᵈ _ _ _ _ _ H lemma infinite_of_not_bdd_below : ¬ bdd_below s → s.infinite := begin diff --git a/src/data/set/function.lean b/src/data/set/function.lean index 40a99087faaa1..7497ee422bfcb 100644 --- a/src/data/set/function.lean +++ b/src/data/set/function.lean @@ -895,7 +895,7 @@ lemma piecewise_le {δ : α → Type*} [Π i, preorder (δ i)] {s : set α} [Π lemma le_piecewise {δ : α → Type*} [Π i, preorder (δ i)] {s : set α} [Π j, decidable (j ∈ s)] {f₁ f₂ g : Π i, δ i} (h₁ : ∀ i ∈ s, g i ≤ f₁ i) (h₂ : ∀ i ∉ s, g i ≤ f₂ i) : g ≤ s.piecewise f₁ f₂ := -@piecewise_le α (λ i, order_dual (δ i)) _ s _ _ _ _ h₁ h₂ +@piecewise_le α (λ i, (δ i)ᵒᵈ) _ s _ _ _ _ h₁ h₂ lemma piecewise_le_piecewise {δ : α → Type*} [Π i, preorder (δ i)] {s : set α} [Π j, decidable (j ∈ s)] {f₁ f₂ g₁ g₂ : Π i, δ i} (h₁ : ∀ i ∈ s, f₁ i ≤ g₁ i) @@ -1015,7 +1015,7 @@ lemma strict_mono_on.inj_on [linear_order α] [preorder β] {f : α → β} {s : lemma strict_anti_on.inj_on [linear_order α] [preorder β] {f : α → β} {s : set α} (H : strict_anti_on f s) : s.inj_on f := -@strict_mono_on.inj_on α (order_dual β) _ _ f s H +@strict_mono_on.inj_on α βᵒᵈ _ _ f s H lemma strict_mono_on.comp [preorder α] [preorder β] [preorder γ] {g : β → γ} {f : α → β} {s : set α} {t : set β} (hg : strict_mono_on g t) diff --git a/src/data/set/intervals/basic.lean b/src/data/set/intervals/basic.lean index c4a7a2815a8a5..2db4ebb1c677c 100644 --- a/src/data/set/intervals/basic.lean +++ b/src/data/set/intervals/basic.lean @@ -186,8 +186,7 @@ Ioo_eq_empty h.not_lt lemma Ici_subset_Ici : Ici a ⊆ Ici b ↔ b ≤ a := ⟨λ h, h $ left_mem_Ici, λ h x hx, h.trans hx⟩ -lemma Iic_subset_Iic : Iic a ⊆ Iic b ↔ a ≤ b := -@Ici_subset_Ici (order_dual α) _ _ _ +lemma Iic_subset_Iic : Iic a ⊆ Iic b ↔ a ≤ b := @Ici_subset_Ici αᵒᵈ _ _ _ lemma Ici_subset_Ioi : Ici a ⊆ Ioi b ↔ b < a := ⟨λ h, h left_mem_Ici, λ h x hx, h.trans_le hx⟩ @@ -485,7 +484,7 @@ classical.by_cases lemma mem_Iic_Iio_of_subset_of_subset {s : set α} (ho : Iio a ⊆ s) (hc : s ⊆ Iic a) : s ∈ ({Iic a, Iio a} : set (set α)) := -@mem_Ici_Ioi_of_subset_of_subset (order_dual α) _ a s ho hc +@mem_Ici_Ioi_of_subset_of_subset αᵒᵈ _ a s ho hc lemma mem_Icc_Ico_Ioc_Ioo_of_subset_of_subset {s : set α} (ho : Ioo a b ⊆ s) (hc : s ⊆ Icc a b) : s ∈ ({Icc a b, Ico a b, Ioc a b, Ioo a b} : set (set α)) := @@ -628,7 +627,7 @@ lemma Ico_subset_Ico_iff (h₁ : a₁ < b₁) : lemma Ioc_subset_Ioc_iff (h₁ : a₁ < b₁) : Ioc a₁ b₁ ⊆ Ioc a₂ b₂ ↔ b₁ ≤ b₂ ∧ a₂ ≤ a₁ := -by { convert @Ico_subset_Ico_iff (order_dual α) _ b₁ b₂ a₁ a₂ h₁; exact (@dual_Ico α _ _ _).symm } +by { convert @Ico_subset_Ico_iff αᵒᵈ _ b₁ b₂ a₁ a₂ h₁; exact (@dual_Ico α _ _ _).symm } lemma Ioo_subset_Ioo_iff [densely_ordered α] (h₁ : a₁ < b₁) : Ioo a₁ b₁ ⊆ Ioo a₂ b₂ ↔ a₂ ≤ a₁ ∧ b₁ ≤ b₂ := diff --git a/src/data/set/intervals/infinite.lean b/src/data/set/intervals/infinite.lean index 70b92219d666e..cdca43d3557a7 100644 --- a/src/data/set/intervals/infinite.lean +++ b/src/data/set/intervals/infinite.lean @@ -62,8 +62,7 @@ section unbounded_above variables [no_max_order α] -lemma Ioi.infinite {a : α} : infinite (Ioi a) := -by apply @Iio.infinite (order_dual α) +lemma Ioi.infinite {a : α} : infinite (Ioi a) := @Iio.infinite αᵒᵈ _ _ _ lemma Ici.infinite {a : α} : infinite (Ici a) := Ioi.infinite.mono Ioi_subset_Ici_self diff --git a/src/data/set/intervals/pi.lean b/src/data/set/intervals/pi.lean index 4ded6da0e95fd..575d49afb813c 100644 --- a/src/data/set/intervals/pi.lean +++ b/src/data/set/intervals/pi.lean @@ -51,7 +51,7 @@ lemma pi_univ_Ioi_subset : pi univ (λ i, Ioi (x i)) ⊆ Ioi x := λ h, nonempty.elim ‹nonempty ι› $ λ i, (h i).not_lt (hz i trivial)⟩ lemma pi_univ_Iio_subset : pi univ (λ i, Iio (x i)) ⊆ Iio x := -@pi_univ_Ioi_subset ι (λ i, order_dual (α i)) _ x _ +@pi_univ_Ioi_subset ι (λ i, (α i)ᵒᵈ) _ x _ lemma pi_univ_Ioo_subset : pi univ (λ i, Ioo (x i) (y i)) ⊆ Ioo x y := λ x hx, ⟨pi_univ_Ioi_subset _ $ λ i hi, (hx i hi).1, pi_univ_Iio_subset _ $ λ i hi, (hx i hi).2⟩ diff --git a/src/data/set/intervals/with_bot_top.lean b/src/data/set/intervals/with_bot_top.lean index 04f6644b2f3de..6f4c2ee095456 100644 --- a/src/data/set/intervals/with_bot_top.lean +++ b/src/data/set/intervals/with_bot_top.lean @@ -95,11 +95,11 @@ end with_top namespace with_bot @[simp] lemma preimage_coe_bot : (coe : α → with_bot α) ⁻¹' {⊥} = (∅ : set α) := -@with_top.preimage_coe_top (order_dual α) +@with_top.preimage_coe_top αᵒᵈ variables [partial_order α] {a b : α} -lemma range_coe : range (coe : α → with_bot α) = Ioi ⊥ := @with_top.range_coe (order_dual α) _ +lemma range_coe : range (coe : α → with_bot α) = Ioi ⊥ := @with_top.range_coe αᵒᵈ _ @[simp] lemma preimage_coe_Ioi : (coe : α → with_bot α) ⁻¹' (Ioi a) = Ioi a := ext $ λ x, coe_lt_coe @[simp] lemma preimage_coe_Ici : (coe : α → with_bot α) ⁻¹' (Ici a) = Ici a := ext $ λ x, coe_le_coe diff --git a/src/data/sum/order.lean b/src/data/sum/order.lean index f6f26e9e4637e..c6324afa960ac 100644 --- a/src/data/sum/order.lean +++ b/src/data/sum/order.lean @@ -438,8 +438,7 @@ rfl @[simp] lemma sum_assoc_symm_apply_inr_inr : (sum_assoc α β γ).symm (inr (inr c)) = inr c := rfl /-- `order_dual` is distributive over `⊕` up to an order isomorphism. -/ -def sum_dual_distrib (α β : Type*) [has_le α] [has_le β] : - order_dual (α ⊕ β) ≃o order_dual α ⊕ order_dual β := +def sum_dual_distrib (α β : Type*) [has_le α] [has_le β] : (α ⊕ β)ᵒᵈ ≃o αᵒᵈ ⊕ βᵒᵈ := { map_rel_iff' := begin rintro (a | a) (b | b), { change inl (to_dual a) ≤ inl (to_dual b) ↔ to_dual (inl a) ≤ to_dual (inl b), @@ -501,8 +500,7 @@ def sum_lex_assoc (α β γ : Type*) [has_le α] [has_le β] [has_le γ] : (α (sum_lex_assoc α β γ).symm (inr (inr c)) = inr c := rfl /-- `order_dual` is antidistributive over `⊕ₗ` up to an order isomorphism. -/ -def sum_lex_dual_antidistrib (α β : Type*) [has_le α] [has_le β] : - order_dual (α ⊕ₗ β) ≃o order_dual β ⊕ₗ order_dual α := +def sum_lex_dual_antidistrib (α β : Type*) [has_le α] [has_le β] : (α ⊕ₗ β)ᵒᵈ ≃o βᵒᵈ ⊕ₗ αᵒᵈ := { map_rel_iff' := begin rintro (a | a) (b | b), simp, { change to_lex (inr $ to_dual a) ≤ to_lex (inr $ to_dual b) ↔ diff --git a/src/field_theory/galois.lean b/src/field_theory/galois.lean index 22f5c16445623..6447db5cb2dac 100644 --- a/src/field_theory/galois.lean +++ b/src/field_theory/galois.lean @@ -251,7 +251,7 @@ by conv { to_rhs, rw [←fixed_field_fixing_subgroup K, /-- The Galois correspondence from intermediate fields to subgroups -/ def intermediate_field_equiv_subgroup [finite_dimensional F E] [is_galois F E] : - intermediate_field F E ≃o order_dual (subgroup (E ≃ₐ[F] E)) := + intermediate_field F E ≃o (subgroup (E ≃ₐ[F] E))ᵒᵈ := { to_fun := intermediate_field.fixing_subgroup, inv_fun := intermediate_field.fixed_field, left_inv := λ K, fixed_field_fixing_subgroup K, diff --git a/src/linear_algebra/affine_space/ordered.lean b/src/linear_algebra/affine_space/ordered.lean index 4fea2a3879290..833ef73fbcd54 100644 --- a/src/linear_algebra/affine_space/ordered.lean +++ b/src/linear_algebra/affine_space/ordered.lean @@ -96,13 +96,13 @@ lemma left_lt_line_map_iff_lt (h : 0 < r) : a < line_map a b r ↔ a < b := iff.trans (by rw line_map_apply_zero) (line_map_lt_line_map_iff_of_lt h) lemma line_map_lt_left_iff_lt (h : 0 < r) : line_map a b r < a ↔ b < a := -@left_lt_line_map_iff_lt k (order_dual E) _ _ _ _ _ _ _ h +@left_lt_line_map_iff_lt k Eᵒᵈ _ _ _ _ _ _ _ h lemma line_map_lt_right_iff_lt (h : r < 1) : line_map a b r < b ↔ a < b := iff.trans (by rw line_map_apply_one) (line_map_lt_line_map_iff_of_lt h) lemma right_lt_line_map_iff_lt (h : r < 1) : b < line_map a b r ↔ b < a := -@line_map_lt_right_iff_lt k (order_dual E) _ _ _ _ _ _ _ h +@line_map_lt_right_iff_lt k Eᵒᵈ _ _ _ _ _ _ _ h end ordered_ring @@ -143,7 +143,7 @@ iff.trans (by rw line_map_apply_zero) (line_map_le_line_map_iff_of_lt h) left_le_line_map_iff_le $ inv_pos.2 zero_lt_two lemma line_map_le_left_iff_le (h : 0 < r) : line_map a b r ≤ a ↔ b ≤ a := -@left_le_line_map_iff_le k (order_dual E) _ _ _ _ _ _ _ h +@left_le_line_map_iff_le k Eᵒᵈ _ _ _ _ _ _ _ h @[simp] lemma midpoint_le_left : midpoint k a b ≤ a ↔ b ≤ a := line_map_le_left_iff_le $ inv_pos.2 zero_lt_two @@ -155,7 +155,7 @@ iff.trans (by rw line_map_apply_one) (line_map_le_line_map_iff_of_lt h) line_map_le_right_iff_le $ inv_lt_one one_lt_two lemma right_le_line_map_iff_le (h : r < 1) : b ≤ line_map a b r ↔ b ≤ a := -@line_map_le_right_iff_le k (order_dual E) _ _ _ _ _ _ _ h +@line_map_le_right_iff_le k Eᵒᵈ _ _ _ _ _ _ _ h @[simp] lemma right_le_midpoint : b ≤ midpoint k a b ↔ b ≤ a := right_le_line_map_iff_le $ inv_lt_one one_lt_two @@ -213,7 +213,7 @@ end segment `[(a, f a), (b, f b)]` if and only if `slope f a b ≤ slope f a c`. -/ lemma line_map_le_map_iff_slope_le_slope_left (h : 0 < r * (b - a)) : line_map (f a) (f b) r ≤ f c ↔ slope f a b ≤ slope f a c := -@map_le_line_map_iff_slope_le_slope_left k (order_dual E) _ _ _ _ f a b r h +@map_le_line_map_iff_slope_le_slope_left k Eᵒᵈ _ _ _ _ f a b r h /-- Given `c = line_map a b r`, `a < c`, the point `(c, f c)` is strictly below the segment `[(a, f a), (b, f b)]` if and only if `slope f a c < slope f a b`. -/ @@ -226,7 +226,7 @@ lt_iff_lt_of_le_iff_le' (line_map_le_map_iff_slope_le_slope_left h) segment `[(a, f a), (b, f b)]` if and only if `slope f a b < slope f a c`. -/ lemma line_map_lt_map_iff_slope_lt_slope_left (h : 0 < r * (b - a)) : line_map (f a) (f b) r < f c ↔ slope f a b < slope f a c := -@map_lt_line_map_iff_slope_lt_slope_left k (order_dual E) _ _ _ _ f a b r h +@map_lt_line_map_iff_slope_lt_slope_left k Eᵒᵈ _ _ _ _ f a b r h /-- Given `c = line_map a b r`, `c < b`, the point `(c, f c)` is non-strictly below the segment `[(a, f a), (b, f b)]` if and only if `slope f a b ≤ slope f c b`. -/ @@ -247,7 +247,7 @@ end segment `[(a, f a), (b, f b)]` if and only if `slope f c b ≤ slope f a b`. -/ lemma line_map_le_map_iff_slope_le_slope_right (h : 0 < (1 - r) * (b - a)) : line_map (f a) (f b) r ≤ f c ↔ slope f c b ≤ slope f a b := -@map_le_line_map_iff_slope_le_slope_right k (order_dual E) _ _ _ _ f a b r h +@map_le_line_map_iff_slope_le_slope_right k Eᵒᵈ _ _ _ _ f a b r h /-- Given `c = line_map a b r`, `c < b`, the point `(c, f c)` is strictly below the segment `[(a, f a), (b, f b)]` if and only if `slope f a b < slope f c b`. -/ @@ -260,7 +260,7 @@ lt_iff_lt_of_le_iff_le' (line_map_le_map_iff_slope_le_slope_right h) segment `[(a, f a), (b, f b)]` if and only if `slope f c b < slope f a b`. -/ lemma line_map_lt_map_iff_slope_lt_slope_right (h : 0 < (1 - r) * (b - a)) : line_map (f a) (f b) r < f c ↔ slope f c b < slope f a b := -@map_lt_line_map_iff_slope_lt_slope_right k (order_dual E) _ _ _ _ f a b r h +@map_lt_line_map_iff_slope_lt_slope_right k Eᵒᵈ _ _ _ _ f a b r h /-- Given `c = line_map a b r`, `a < c < b`, the point `(c, f c)` is non-strictly below the segment `[(a, f a), (b, f b)]` if and only if `slope f a c ≤ slope f c b`. -/ @@ -277,7 +277,7 @@ end segment `[(a, f a), (b, f b)]` if and only if `slope f c b ≤ slope f a c`. -/ lemma line_map_le_map_iff_slope_le_slope (hab : a < b) (h₀ : 0 < r) (h₁ : r < 1) : line_map (f a) (f b) r ≤ f c ↔ slope f c b ≤ slope f a c := -@map_le_line_map_iff_slope_le_slope k (order_dual E) _ _ _ _ _ _ _ _ hab h₀ h₁ +@map_le_line_map_iff_slope_le_slope k Eᵒᵈ _ _ _ _ _ _ _ _ hab h₀ h₁ /-- Given `c = line_map a b r`, `a < c < b`, the point `(c, f c)` is strictly below the segment `[(a, f a), (b, f b)]` if and only if `slope f a c < slope f c b`. -/ @@ -290,6 +290,6 @@ lt_iff_lt_of_le_iff_le' (line_map_le_map_iff_slope_le_slope hab h₀ h₁) segment `[(a, f a), (b, f b)]` if and only if `slope f c b < slope f a c`. -/ lemma line_map_lt_map_iff_slope_lt_slope (hab : a < b) (h₀ : 0 < r) (h₁ : r < 1) : line_map (f a) (f b) r < f c ↔ slope f c b < slope f a c := -@map_lt_line_map_iff_slope_lt_slope k (order_dual E) _ _ _ _ _ _ _ _ hab h₀ h₁ +@map_lt_line_map_iff_slope_lt_slope k Eᵒᵈ _ _ _ _ _ _ _ _ hab h₀ h₁ end linear_ordered_field diff --git a/src/linear_algebra/basic.lean b/src/linear_algebra/basic.lean index 43fde04ad6c3d..558762474a163 100644 --- a/src/linear_algebra/basic.lean +++ b/src/linear_algebra/basic.lean @@ -971,7 +971,7 @@ end The decreasing sequence of submodules consisting of the ranges of the iterates of a linear map. -/ @[simps] -def iterate_range (f : M →ₗ[R] M) : ℕ →o order_dual (submodule R M) := +def iterate_range (f : M →ₗ[R] M) : ℕ →o (submodule R M)ᵒᵈ := ⟨λ n, (f ^ n).range, λ n m w x h, begin obtain ⟨c, rfl⟩ := le_iff_exists_add.mp w, rw linear_map.mem_range at h, diff --git a/src/linear_algebra/prod.lean b/src/linear_algebra/prod.lean index b2ca23c5c34c1..a2898792568a7 100644 --- a/src/linear_algebra/prod.lean +++ b/src/linear_algebra/prod.lean @@ -686,7 +686,7 @@ noncomputable def tunnel' (f : M × N →ₗ[R] M) (i : injective f) : Give an injective map `f : M × N →ₗ[R] M` we can find a nested sequence of submodules all isomorphic to `M`. -/ -def tunnel (f : M × N →ₗ[R] M) (i : injective f) : ℕ →o order_dual (submodule R M) := +def tunnel (f : M × N →ₗ[R] M) (i : injective f) : ℕ →o (submodule R M)ᵒᵈ := ⟨λ n, (tunnel' f i n).1, monotone_nat_of_le_succ (λ n, begin dsimp [tunnel', tunnel_aux], rw [submodule.map_comp, submodule.map_comp], diff --git a/src/measure_theory/constructions/borel_space.lean b/src/measure_theory/constructions/borel_space.lean index d00dd27457ff2..fb54f747b226f 100644 --- a/src/measure_theory/constructions/borel_space.lean +++ b/src/measure_theory/constructions/borel_space.lean @@ -147,7 +147,7 @@ begin end lemma borel_eq_generate_from_Ioi : borel α = generate_from (range Ioi) := -@borel_eq_generate_from_Iio (order_dual α) _ (by apply_instance : second_countable_topology α) _ _ +@borel_eq_generate_from_Iio αᵒᵈ _ (by apply_instance : second_countable_topology α) _ _ end order_topology @@ -232,13 +232,13 @@ end tactic @[priority 100] instance order_dual.opens_measurable_space {α : Type*} [topological_space α] [measurable_space α] [h : opens_measurable_space α] : - opens_measurable_space (order_dual α) := + opens_measurable_space αᵒᵈ := { borel_le := h.borel_le } @[priority 100] instance order_dual.borel_space {α : Type*} [topological_space α] [measurable_space α] [h : borel_space α] : - borel_space (order_dual α) := + borel_space αᵒᵈ := { measurable_eq := h.measurable_eq } /-- In a `borel_space` all open sets are measurable. -/ @@ -619,7 +619,7 @@ lemma ext_of_Ioc_finite {α : Type*} [topological_space α] {m : measurable_spac [borel_space α] (μ ν : measure α) [is_finite_measure μ] (hμν : μ univ = ν univ) (h : ∀ ⦃a b⦄, a < b → μ (Ioc a b) = ν (Ioc a b)) : μ = ν := begin - refine @ext_of_Ico_finite (order_dual α) _ _ _ _ _ ‹_› μ ν _ hμν (λ a b hab, _), + refine @ext_of_Ico_finite αᵒᵈ _ _ _ _ _ ‹_› μ ν _ hμν (λ a b hab, _), erw dual_Ico, exact h hab end @@ -655,7 +655,7 @@ lemma ext_of_Ioc' {α : Type*} [topological_space α] {m : measurable_space α} [no_min_order α] (μ ν : measure α) (hμ : ∀ ⦃a b⦄, a < b → μ (Ioc a b) ≠ ∞) (h : ∀ ⦃a b⦄, a < b → μ (Ioc a b) = ν (Ioc a b)) : μ = ν := begin - refine @ext_of_Ico' (order_dual α) _ _ _ _ _ ‹_› _ μ ν _ _; + refine @ext_of_Ico' αᵒᵈ _ _ _ _ _ ‹_› _ μ ν _ _; intros a b hab; erw dual_Ico, exacts [hμ hab, h hab] end @@ -697,7 +697,7 @@ intervals. -/ lemma ext_of_Ici {α : Type*} [topological_space α] {m : measurable_space α} [second_countable_topology α] [linear_order α] [order_topology α] [borel_space α] (μ ν : measure α) [is_finite_measure μ] (h : ∀ a, μ (Ici a) = ν (Ici a)) : μ = ν := -@ext_of_Iic (order_dual α) _ _ _ _ _ ‹_› _ _ _ h +@ext_of_Iic αᵒᵈ _ _ _ _ _ ‹_› _ _ _ h end measure_theory.measure @@ -1129,12 +1129,12 @@ ae_measurable_restrict_of_measurable_subtype hs this.measurable protected lemma antitone.measurable [linear_order β] [order_closed_topology β] {f : β → α} (hf : antitone f) : measurable f := -@monotone.measurable (order_dual α) β _ _ ‹_› _ _ _ _ _ ‹_› _ _ _ hf +@monotone.measurable αᵒᵈ β _ _ ‹_› _ _ _ _ _ ‹_› _ _ _ hf lemma ae_measurable_restrict_of_antitone_on [linear_order β] [order_closed_topology β] {μ : measure β} {s : set β} (hs : measurable_set s) {f : β → α} (hf : antitone_on f s) : ae_measurable f (μ.restrict s) := -@ae_measurable_restrict_of_monotone_on (order_dual α) β _ _ ‹_› _ _ _ _ _ ‹_› _ _ _ _ hs _ hf +@ae_measurable_restrict_of_monotone_on αᵒᵈ β _ _ ‹_› _ _ _ _ _ ‹_› _ _ _ _ hs _ hf end linear_order diff --git a/src/measure_theory/function/ess_sup.lean b/src/measure_theory/function/ess_sup.lean index 5c3d3879924d6..469004444c966 100644 --- a/src/measure_theory/function/ess_sup.lean +++ b/src/measure_theory/function/ess_sup.lean @@ -46,7 +46,7 @@ lemma ess_sup_congr_ae {f g : α → β} (hfg : f =ᵐ[μ] g) : ess_sup f μ = e limsup_congr hfg lemma ess_inf_congr_ae {f g : α → β} (hfg : f =ᵐ[μ] g) : ess_inf f μ = ess_inf g μ := -@ess_sup_congr_ae α (order_dual β) _ _ _ _ _ hfg +@ess_sup_congr_ae α βᵒᵈ _ _ _ _ _ hfg end conditionally_complete_lattice @@ -73,7 +73,7 @@ le_bot_iff.mp (Inf_le (by simp [set.mem_set_of_eq, eventually_le, ae_iff])) @[simp] lemma ess_inf_measure_zero {m : measurable_space α} {f : α → β} : ess_inf f (0 : measure α) = ⊤ := -@ess_sup_measure_zero α (order_dual β) _ _ _ +@ess_sup_measure_zero α βᵒᵈ _ _ _ lemma ess_sup_mono_ae {f g : α → β} (hfg : f ≤ᵐ[μ] g) : ess_sup f μ ≤ ess_sup g μ := limsup_le_limsup hfg @@ -96,10 +96,10 @@ begin end lemma ess_inf_const (c : β) (hμ : μ ≠ 0) : ess_inf (λ x : α, c) μ = c := -@ess_sup_const α (order_dual β) _ _ _ _ hμ +@ess_sup_const α βᵒᵈ _ _ _ _ hμ lemma le_ess_inf_of_ae_le {f : α → β} (c : β) (hf : (λ _, c) ≤ᵐ[μ] f) : c ≤ ess_inf f μ := -@ess_sup_le_of_ae_le α (order_dual β) _ _ _ _ c hf +@ess_sup_le_of_ae_le α βᵒᵈ _ _ _ _ c hf lemma ess_sup_const_bot : ess_sup (λ x : α, (⊥ : β)) μ = (⊥ : β) := limsup_const_bot @@ -118,7 +118,7 @@ end lemma order_iso.ess_inf_apply {m : measurable_space α} {γ} [complete_lattice γ] (f : α → β) (μ : measure α) (g : β ≃o γ) : g (ess_inf f μ) = ess_inf (λ x, g (f x)) μ := -@order_iso.ess_sup_apply α (order_dual β) _ _ (order_dual γ) _ _ _ g.dual +@order_iso.ess_sup_apply α βᵒᵈ _ _ γᵒᵈ _ _ _ g.dual lemma ess_sup_mono_measure {f : α → β} (hμν : ν ≪ μ) : ess_sup f ν ≤ ess_sup f μ := begin @@ -207,7 +207,7 @@ lemma ae_lt_of_ess_sup_lt {f : α → β} {x : β} (hf : ess_sup f μ < x) : ∀ filter.eventually_lt_of_limsup_lt hf lemma ae_lt_of_lt_ess_inf {f : α → β} {x : β} (hf : x < ess_inf f μ) : ∀ᵐ y ∂μ, x < f y := -@ae_lt_of_ess_sup_lt α (order_dual β) _ _ _ _ _ hf +@ae_lt_of_ess_sup_lt α βᵒᵈ _ _ _ _ _ hf lemma ess_sup_indicator_eq_ess_sup_restrict [has_zero β] {s : set α} {f : α → β} (hf : 0 ≤ᵐ[μ.restrict s] f) (hs : measurable_set s) (hs_not_null : μ s ≠ 0) : diff --git a/src/measure_theory/function/locally_integrable.lean b/src/measure_theory/function/locally_integrable.lean index b2cd7339b04e7..e8d07e8aeebf3 100644 --- a/src/measure_theory/function/locally_integrable.lean +++ b/src/measure_theory/function/locally_integrable.lean @@ -180,12 +180,12 @@ end lemma antitone_on.integrable_on_compact (hs : is_compact s) (hanti : antitone_on f s) : integrable_on f s μ := -@monotone_on.integrable_on_compact X (order_dual E) _ _ _ _ _ _ _ _ _ _ _ _ _ _ hs hanti +hanti.dual_right.integrable_on_compact hs lemma monotone.locally_integrable (hmono : monotone f) : locally_integrable f μ := -λ s hs, monotone_on.integrable_on_compact hs (λ x y _ _ hxy, hmono hxy) +λ s hs, (hmono.monotone_on _).integrable_on_compact hs lemma antitone.locally_integrable (hanti : antitone f) : locally_integrable f μ := -@monotone.locally_integrable X (order_dual E) _ _ _ _ _ _ _ _ _ _ _ _ _ hanti +hanti.dual_right.locally_integrable end monotone diff --git a/src/measure_theory/integral/interval_integral.lean b/src/measure_theory/integral/interval_integral.lean index bdd9626da14f9..af85dc5380c35 100644 --- a/src/measure_theory/integral/interval_integral.lean +++ b/src/measure_theory/integral/interval_integral.lean @@ -358,7 +358,7 @@ end lemma antitone_on.interval_integrable {u : ℝ → E} {a b : ℝ} (hu : antitone_on u (interval a b)) : interval_integrable u μ a b := -@monotone_on.interval_integrable (order_dual E) _ _ _ _ _ _ _ _ _ hu +hu.dual_right.interval_integrable lemma monotone.interval_integrable {u : ℝ → E} {a b : ℝ} (hu : monotone u) : interval_integrable u μ a b := diff --git a/src/measure_theory/lattice.lean b/src/measure_theory/lattice.lean index 069c10d0c9adb..87b9da7e7679c 100644 --- a/src/measure_theory/lattice.lean +++ b/src/measure_theory/lattice.lean @@ -63,23 +63,19 @@ variables {M : Type*} [measurable_space M] section order_dual @[priority 100] -instance order_dual.has_measurable_sup [has_inf M] [has_measurable_inf M] : - has_measurable_sup (order_dual M) := +instance [has_inf M] [has_measurable_inf M] : has_measurable_sup Mᵒᵈ := ⟨@measurable_const_inf M _ _ _, @measurable_inf_const M _ _ _⟩ @[priority 100] -instance order_dual.has_measurable_inf [has_sup M] [has_measurable_sup M] : - has_measurable_inf (order_dual M) := +instance [has_sup M] [has_measurable_sup M] : has_measurable_inf Mᵒᵈ := ⟨@measurable_const_sup M _ _ _, @measurable_sup_const M _ _ _⟩ @[priority 100] -instance order_dual.has_measurable_sup₂ [has_inf M] [has_measurable_inf₂ M] : - has_measurable_sup₂ (order_dual M) := +instance [has_inf M] [has_measurable_inf₂ M] : has_measurable_sup₂ Mᵒᵈ := ⟨@measurable_inf M _ _ _⟩ @[priority 100] -instance order_dual.has_measurable_inf₂ [has_sup M] [has_measurable_sup₂ M] : - has_measurable_inf₂ (order_dual M) := +instance [has_sup M] [has_measurable_sup₂ M] : has_measurable_inf₂ Mᵒᵈ := ⟨@measurable_sup M _ _ _⟩ end order_dual diff --git a/src/measure_theory/measurable_space_def.lean b/src/measure_theory/measurable_space_def.lean index c136d5f854c29..c7cf0139715b6 100644 --- a/src/measure_theory/measurable_space_def.lean +++ b/src/measure_theory/measurable_space_def.lean @@ -55,7 +55,7 @@ structure measurable_space (α : Type*) := attribute [class] measurable_space -instance [h : measurable_space α] : measurable_space (order_dual α) := h +instance [h : measurable_space α] : measurable_space αᵒᵈ := h section diff --git a/src/measure_theory/measure/measure_space.lean b/src/measure_theory/measure/measure_space.lean index 2cf0c911f4c6a..f86260b4fabfc 100644 --- a/src/measure_theory/measure/measure_space.lean +++ b/src/measure_theory/measure/measure_space.lean @@ -2046,8 +2046,7 @@ variables [partial_order α] {a b : α} lemma Iio_ae_eq_Iic' (ha : μ {a} = 0) : Iio a =ᵐ[μ] Iic a := by rw [←Iic_diff_right, diff_ae_eq_self, measure_mono_null (set.inter_subset_right _ _) ha] -lemma Ioi_ae_eq_Ici' (ha : μ {a} = 0) : Ioi a =ᵐ[μ] Ici a := -@Iio_ae_eq_Iic' (order_dual α) ‹_› ‹_› _ _ ha +lemma Ioi_ae_eq_Ici' (ha : μ {a} = 0) : Ioi a =ᵐ[μ] Ici a := @Iio_ae_eq_Iic' αᵒᵈ ‹_› ‹_› _ _ ha lemma Ioo_ae_eq_Ioc' (hb : μ {b} = 0) : Ioo a b =ᵐ[μ] Ioc a b := (ae_eq_refl _).inter (Iio_ae_eq_Iic' hb) diff --git a/src/measure_theory/pi_system.lean b/src/measure_theory/pi_system.lean index 485382125312d..efaaa9aa0f3cb 100644 --- a/src/measure_theory/pi_system.lean +++ b/src/measure_theory/pi_system.lean @@ -87,7 +87,7 @@ lemma is_pi_system_Iio : is_pi_system (range Iio : set (set α)) := @image_univ α _ Iio ▸ is_pi_system_image_Iio univ lemma is_pi_system_image_Ioi (s : set α) : is_pi_system (Ioi '' s) := -@is_pi_system_image_Iio (order_dual α) _ s +@is_pi_system_image_Iio αᵒᵈ _ s lemma is_pi_system_Ioi : is_pi_system (range Ioi : set (set α)) := @image_univ α _ Ioi ▸ is_pi_system_image_Ioi univ diff --git a/src/order/antisymmetrization.lean b/src/order/antisymmetrization.lean index c6ca0a435e52e..4e8098fb5a811 100644 --- a/src/order/antisymmetrization.lean +++ b/src/order/antisymmetrization.lean @@ -169,7 +169,7 @@ variables (α) /-- `antisymmetrization` and `order_dual` commute. -/ def order_iso.dual_antisymmetrization : - order_dual (antisymmetrization α (≤)) ≃o antisymmetrization (order_dual α) (≤) := + (antisymmetrization α (≤))ᵒᵈ ≃o antisymmetrization αᵒᵈ (≤) := { to_fun := quotient.map' id $ λ _ _, and.symm, inv_fun := quotient.map' id $ λ _ _, and.symm, left_inv := λ a, quotient.induction_on' a $ λ a, by simp_rw [quotient.map'_mk', id], diff --git a/src/order/atoms.lean b/src/order/atoms.lean index 052dcd0a24a51..bce44154fb480 100644 --- a/src/order/atoms.lean +++ b/src/order/atoms.lean @@ -155,20 +155,17 @@ export is_atomic (eq_bot_or_exists_atom_le) is_coatomic (eq_top_or_exists_le_coa variable {α} -@[simp] theorem is_coatomic_dual_iff_is_atomic [order_bot α] : - is_coatomic (order_dual α) ↔ is_atomic α := +@[simp] lemma is_coatomic_dual_iff_is_atomic [order_bot α] : is_coatomic αᵒᵈ ↔ is_atomic α := ⟨λ h, ⟨λ b, by apply h.eq_top_or_exists_le_coatom⟩, λ h, ⟨λ b, by apply h.eq_bot_or_exists_atom_le⟩⟩ -@[simp] theorem is_atomic_dual_iff_is_coatomic [order_top α] : - is_atomic (order_dual α) ↔ is_coatomic α := +@[simp] lemma is_atomic_dual_iff_is_coatomic [order_top α] : is_atomic αᵒᵈ ↔ is_coatomic α := ⟨λ h, ⟨λ b, by apply h.eq_bot_or_exists_atom_le⟩, λ h, ⟨λ b, by apply h.eq_top_or_exists_le_coatom⟩⟩ namespace is_atomic variables [order_bot α] [is_atomic α] -instance is_coatomic_dual : is_coatomic (order_dual α) := -is_coatomic_dual_iff_is_atomic.2 ‹is_atomic α› +instance is_coatomic_dual : is_coatomic αᵒᵈ := is_coatomic_dual_iff_is_atomic.2 ‹is_atomic α› instance {x : α} : is_atomic (set.Iic x) := ⟨λ ⟨y, hy⟩, (eq_bot_or_exists_atom_le y).imp subtype.mk_eq_mk.2 @@ -180,8 +177,7 @@ namespace is_coatomic variables [order_top α] [is_coatomic α] -instance is_coatomic : is_atomic (order_dual α) := -is_atomic_dual_iff_is_coatomic.2 ‹is_coatomic α› +instance is_coatomic : is_atomic αᵒᵈ := is_atomic_dual_iff_is_coatomic.2 ‹is_coatomic α› instance {x : α} : is_coatomic (set.Ici x) := ⟨λ ⟨y, hy⟩, (eq_top_or_exists_le_coatom y).imp subtype.mk_eq_mk.2 @@ -219,16 +215,16 @@ export is_atomistic (eq_Sup_atoms) is_coatomistic (eq_Inf_coatoms) variable {α} @[simp] -theorem is_coatomistic_dual_iff_is_atomistic : is_coatomistic (order_dual α) ↔ is_atomistic α := +theorem is_coatomistic_dual_iff_is_atomistic : is_coatomistic αᵒᵈ ↔ is_atomistic α := ⟨λ h, ⟨λ b, by apply h.eq_Inf_coatoms⟩, λ h, ⟨λ b, by apply h.eq_Sup_atoms⟩⟩ @[simp] -theorem is_atomistic_dual_iff_is_coatomistic : is_atomistic (order_dual α) ↔ is_coatomistic α := +theorem is_atomistic_dual_iff_is_coatomistic : is_atomistic αᵒᵈ ↔ is_coatomistic α := ⟨λ h, ⟨λ b, by apply h.eq_Sup_atoms⟩, λ h, ⟨λ b, by apply h.eq_Inf_coatoms⟩⟩ namespace is_atomistic -instance is_coatomistic_dual [h : is_atomistic α] : is_coatomistic (order_dual α) := +instance is_coatomistic_dual [h : is_atomistic α] : is_coatomistic αᵒᵈ := is_coatomistic_dual_iff_is_atomistic.2 h variable [is_atomistic α] @@ -270,7 +266,7 @@ end is_atomistic namespace is_coatomistic -instance is_atomistic_dual [h : is_coatomistic α] : is_atomistic (order_dual α) := +instance is_atomistic_dual [h : is_coatomistic α] : is_atomistic αᵒᵈ := is_atomistic_dual_iff_is_coatomistic.2 h variable [is_coatomistic α] @@ -292,12 +288,12 @@ class is_simple_order (α : Type*) [has_le α] [bounded_order α] extends nontri export is_simple_order (eq_bot_or_eq_top) theorem is_simple_order_iff_is_simple_order_order_dual [has_le α] [bounded_order α] : - is_simple_order α ↔ is_simple_order (order_dual α) := + is_simple_order α ↔ is_simple_order αᵒᵈ := begin split; intro i; haveI := i, { exact { exists_pair_ne := @exists_pair_ne α _, eq_bot_or_eq_top := λ a, or.symm (eq_bot_or_eq_top ((order_dual.of_dual a)) : _ ∨ _) } }, - { exact { exists_pair_ne := @exists_pair_ne (order_dual α) _, + { exact { exists_pair_ne := @exists_pair_ne αᵒᵈ _, eq_bot_or_eq_top := λ a, or.symm (eq_bot_or_eq_top (order_dual.to_dual a)) } } end @@ -314,7 +310,7 @@ section is_simple_order variables [partial_order α] [bounded_order α] [is_simple_order α] -instance {α} [has_le α] [bounded_order α] [is_simple_order α] : is_simple_order (order_dual α) := +instance {α} [has_le α] [bounded_order α] [is_simple_order α] : is_simple_order αᵒᵈ := is_simple_order_iff_is_simple_order_order_dual.1 (by apply_instance) /-- A simple `bounded_order` induces a preorder. This is not an instance to prevent loops. -/ diff --git a/src/order/basic.lean b/src/order/basic.lean index 48a73f4fa6044..d5ed39b3f15f1 100644 --- a/src/order/basic.lean +++ b/src/order/basic.lean @@ -14,7 +14,7 @@ classes and allows to transfer order instances. ## Type synonyms -* `order_dual α` : A type synonym reversing the meaning of all inequalities. +* `order_dual α` : A type synonym reversing the meaning of all inequalities, with notation `αᵒᵈ`. * `as_linear_order α`: A type synonym to promote `partial_order α` to `linear_order α` using `is_total α (≤)`. @@ -389,37 +389,40 @@ instance order.preimage.decidable {α β} (f : α → β) (s : β → β → Pro /-! ### Order dual -/ -/-- Type synonym to equip a type with the dual order: `≤` means `≥` and `<` means `>`. -/ +/-- Type synonym to equip a type with the dual order: `≤` means `≥` and `<` means `>`. `αᵒᵈ` is +notation for `order_dual α`. -/ def order_dual (α : Type*) : Type* := α +notation α `ᵒᵈ`:std.prec.max_plus := order_dual α + namespace order_dual -instance (α : Type*) [h : nonempty α] : nonempty (order_dual α) := h -instance (α : Type*) [h : subsingleton α] : subsingleton (order_dual α) := h -instance (α : Type*) [has_le α] : has_le (order_dual α) := ⟨λ x y : α, y ≤ x⟩ -instance (α : Type*) [has_lt α] : has_lt (order_dual α) := ⟨λ x y : α, y < x⟩ -instance (α : Type*) [has_zero α] : has_zero (order_dual α) := ⟨(0 : α)⟩ +instance (α : Type*) [h : nonempty α] : nonempty αᵒᵈ := h +instance (α : Type*) [h : subsingleton α] : subsingleton αᵒᵈ := h +instance (α : Type*) [has_le α] : has_le αᵒᵈ := ⟨λ x y : α, y ≤ x⟩ +instance (α : Type*) [has_lt α] : has_lt αᵒᵈ := ⟨λ x y : α, y < x⟩ +instance (α : Type*) [has_zero α] : has_zero αᵒᵈ := ⟨(0 : α)⟩ -- `dual_le` and `dual_lt` should not be simp lemmas: --- they cause a loop since `α` and `order_dual α` are definitionally equal +-- they cause a loop since `α` and `αᵒᵈ` are definitionally equal lemma dual_le [has_le α] {a b : α} : - @has_le.le (order_dual α) _ a b ↔ @has_le.le α _ b a := iff.rfl + @has_le.le αᵒᵈ _ a b ↔ @has_le.le α _ b a := iff.rfl lemma dual_lt [has_lt α] {a b : α} : - @has_lt.lt (order_dual α) _ a b ↔ @has_lt.lt α _ b a := iff.rfl + @has_lt.lt αᵒᵈ _ a b ↔ @has_lt.lt α _ b a := iff.rfl -instance (α : Type*) [preorder α] : preorder (order_dual α) := +instance (α : Type*) [preorder α] : preorder αᵒᵈ := { le_refl := le_refl, le_trans := λ a b c hab hbc, hbc.trans hab, lt_iff_le_not_le := λ _ _, lt_iff_le_not_le, .. order_dual.has_le α, .. order_dual.has_lt α } -instance (α : Type*) [partial_order α] : partial_order (order_dual α) := +instance (α : Type*) [partial_order α] : partial_order αᵒᵈ := { le_antisymm := λ a b hab hba, @le_antisymm α _ a b hba hab, .. order_dual.preorder α } -instance (α : Type*) [linear_order α] : linear_order (order_dual α) := +instance (α : Type*) [linear_order α] : linear_order αᵒᵈ := { le_total := λ a b : α, le_total b a, decidable_le := (infer_instance : decidable_rel (λ a b : α, b ≤ a)), decidable_lt := (infer_instance : decidable_rel (λ a b : α, b < a)), @@ -429,18 +432,18 @@ instance (α : Type*) [linear_order α] : linear_order (order_dual α) := max_def := @linear_order.min_def α _, .. order_dual.partial_order α } -instance : Π [inhabited α], inhabited (order_dual α) := id +instance : Π [inhabited α], inhabited αᵒᵈ := id theorem preorder.dual_dual (α : Type*) [H : preorder α] : - order_dual.preorder (order_dual α) = H := + order_dual.preorder αᵒᵈ = H := preorder.ext $ λ _ _, iff.rfl theorem partial_order.dual_dual (α : Type*) [H : partial_order α] : - order_dual.partial_order (order_dual α) = H := + order_dual.partial_order αᵒᵈ = H := partial_order.ext $ λ _ _, iff.rfl theorem linear_order.dual_dual (α : Type*) [H : linear_order α] : - order_dual.linear_order (order_dual α) = H := + order_dual.linear_order αᵒᵈ = H := linear_order.ext $ λ _ _, iff.rfl end order_dual @@ -626,7 +629,7 @@ lemma exists_between [has_lt α] [densely_ordered α] : densely_ordered.dense instance order_dual.densely_ordered (α : Type u) [has_lt α] [densely_ordered α] : - densely_ordered (order_dual α) := + densely_ordered αᵒᵈ := ⟨λ a₁ a₂ ha, (@exists_between α _ _ _ _ ha).imp $ λ a, and.symm⟩ lemma le_of_forall_le_of_dense [linear_order α] [densely_ordered α] {a₁ a₂ : α} diff --git a/src/order/boolean_algebra.lean b/src/order/boolean_algebra.lean index 85efade06d5a9..beaa2b5a72bc3 100644 --- a/src/order/boolean_algebra.lean +++ b/src/order/boolean_algebra.lean @@ -761,7 +761,7 @@ section boolean_algebra variables [boolean_algebra α] --TODO@Yaël: Once we have co-Heyting algebras, we won't need to go through `boolean_algebra.of_core` -instance : boolean_algebra (order_dual α) := +instance : boolean_algebra αᵒᵈ := boolean_algebra.of_core { compl := λ a, to_dual (of_dual a)ᶜ, inf_compl_le_bot := λ _, sup_compl_eq_top.ge, diff --git a/src/order/bounded.lean b/src/order/bounded.lean index a2ee26dd25ee9..4c29d5138dae0 100644 --- a/src/order/bounded.lean +++ b/src/order/bounded.lean @@ -41,7 +41,7 @@ lemma unbounded_lt_iff [linear_order α] : unbounded (<) s ↔ ∀ a, ∃ b ∈ by simp only [unbounded, not_lt] lemma unbounded_ge_of_forall_exists_gt [preorder α] (h : ∀ a, ∃ b ∈ s, b < a) : unbounded (≥) s := -@unbounded_le_of_forall_exists_lt (order_dual α) _ _ h +@unbounded_le_of_forall_exists_lt αᵒᵈ _ _ h lemma unbounded_ge_iff [linear_order α] : unbounded (≥) s ↔ ∀ a, ∃ b ∈ s, b < a := ⟨λ h a, let ⟨b, hb, hba⟩ := h a in ⟨b, hb, lt_of_not_ge hba⟩, unbounded_ge_of_forall_exists_gt⟩ @@ -90,11 +90,11 @@ lemma unbounded_gt_of_unbounded_ge [preorder α] (h : unbounded (≥) s) : unbou λ a, let ⟨b, hb, hba⟩ := h a in ⟨b, hb, λ hba', hba (le_of_lt hba')⟩ lemma bounded_ge_iff_bounded_gt [preorder α] [no_min_order α] : bounded (≥) s ↔ bounded (>) s := -@bounded_le_iff_bounded_lt (order_dual α) _ _ _ +@bounded_le_iff_bounded_lt αᵒᵈ _ _ _ lemma unbounded_gt_iff_unbounded_ge [preorder α] [no_min_order α] : unbounded (>) s ↔ unbounded (≥) s := -@unbounded_lt_iff_unbounded_le (order_dual α) _ _ _ +@unbounded_lt_iff_unbounded_le αᵒᵈ _ _ _ /-! ### The universal set -/ @@ -296,52 +296,52 @@ end theorem bounded_ge_inter_not_ge [semilattice_inf α] (a : α) : bounded (≥) (s ∩ {b | ¬ a ≤ b}) ↔ bounded (≥) s := -@bounded_le_inter_not_le (order_dual α) s _ a +@bounded_le_inter_not_le αᵒᵈ s _ a theorem unbounded_ge_inter_not_ge [semilattice_inf α] (a : α) : unbounded (≥) (s ∩ {b | ¬ a ≤ b}) ↔ unbounded (≥) s := -@unbounded_le_inter_not_le (order_dual α) s _ a +@unbounded_le_inter_not_le αᵒᵈ s _ a theorem bounded_ge_inter_gt [linear_order α] (a : α) : bounded (≥) (s ∩ {b | b < a}) ↔ bounded (≥) s := -@bounded_le_inter_lt (order_dual α) s _ a +@bounded_le_inter_lt αᵒᵈ s _ a theorem unbounded_ge_inter_gt [linear_order α] (a : α) : unbounded (≥) (s ∩ {b | b < a}) ↔ unbounded (≥) s := -@unbounded_le_inter_lt (order_dual α) s _ a +@unbounded_le_inter_lt αᵒᵈ s _ a theorem bounded_ge_inter_ge [linear_order α] (a : α) : bounded (≥) (s ∩ {b | b ≤ a}) ↔ bounded (≥) s := -@bounded_le_inter_le (order_dual α) s _ a +@bounded_le_inter_le αᵒᵈ s _ a theorem unbounded_ge_iff_unbounded_inter_ge [linear_order α] (a : α) : unbounded (≥) (s ∩ {b | b ≤ a}) ↔ unbounded (≥) s := -@unbounded_le_inter_le (order_dual α) s _ a +@unbounded_le_inter_le αᵒᵈ s _ a /-! #### Greater than -/ theorem bounded_gt_inter_not_gt [semilattice_inf α] (a : α) : bounded (>) (s ∩ {b | ¬ a < b}) ↔ bounded (>) s := -@bounded_lt_inter_not_lt (order_dual α) s _ a +@bounded_lt_inter_not_lt αᵒᵈ s _ a theorem unbounded_gt_inter_not_gt [semilattice_inf α] (a : α) : unbounded (>) (s ∩ {b | ¬ a < b}) ↔ unbounded (>) s := -@unbounded_lt_inter_not_lt (order_dual α) s _ a +@unbounded_lt_inter_not_lt αᵒᵈ s _ a theorem bounded_gt_inter_ge [linear_order α] (a : α) : bounded (>) (s ∩ {b | b ≤ a}) ↔ bounded (>) s := -@bounded_lt_inter_le (order_dual α) s _ a +@bounded_lt_inter_le αᵒᵈ s _ a theorem unbounded_inter_ge [linear_order α] (a : α) : unbounded (>) (s ∩ {b | b ≤ a}) ↔ unbounded (>) s := -@unbounded_lt_inter_le (order_dual α) s _ a +@unbounded_lt_inter_le αᵒᵈ s _ a theorem bounded_gt_inter_gt [linear_order α] [no_min_order α] (a : α) : bounded (>) (s ∩ {b | b < a}) ↔ bounded (>) s := -@bounded_lt_inter_lt (order_dual α) s _ _ a +@bounded_lt_inter_lt αᵒᵈ s _ _ a theorem unbounded_gt_inter_gt [linear_order α] [no_min_order α] (a : α) : unbounded (>) (s ∩ {b | b < a}) ↔ unbounded (>) s := -@unbounded_lt_inter_lt (order_dual α) s _ _ a +@unbounded_lt_inter_lt αᵒᵈ s _ _ a end set diff --git a/src/order/bounded_order.lean b/src/order/bounded_order.lean index dc463a734e37e..1b937d2324ae8 100644 --- a/src/order/bounded_order.lean +++ b/src/order/bounded_order.lean @@ -958,10 +958,10 @@ instance [lattice α] : lattice (with_top α) := { ..with_top.semilattice_sup, ..with_top.semilattice_inf } instance decidable_le [has_le α] [@decidable_rel α (≤)] : @decidable_rel (with_top α) (≤) := -λ x y, @with_bot.decidable_le (order_dual α) _ _ y x +λ x y, @with_bot.decidable_le αᵒᵈ _ _ y x instance decidable_lt [has_lt α] [@decidable_rel α (<)] : @decidable_rel (with_top α) (<) := -λ x y, @with_bot.decidable_lt (order_dual α) _ _ y x +λ x y, @with_bot.decidable_lt αᵒᵈ _ _ y x instance is_total_le [has_le α] [is_total α (≤)] : is_total (with_top α) (≤) := ⟨λ a b, match a, b with @@ -989,11 +989,11 @@ have acc_some : ∀ a : α, acc ((<) : with_top α → with_top α → Prop) (so (λ _ _, acc_some _))) acc_some⟩ lemma well_founded_gt [preorder α] (h : @well_founded α (>)) : @well_founded (with_top α) (>) := -@with_bot.well_founded_lt (order_dual α) _ h +@with_bot.well_founded_lt αᵒᵈ _ h lemma _root_.with_bot.well_founded_gt [preorder α] (h : @well_founded α (>)) : @well_founded (with_bot α) (>) := -@with_top.well_founded_lt (order_dual α) _ h +@with_top.well_founded_lt αᵒᵈ _ h instance [has_lt α] [densely_ordered α] [no_max_order α] : densely_ordered (with_top α) := ⟨ λ a b, @@ -1087,18 +1087,18 @@ end subtype namespace order_dual variable (α) -instance [has_bot α] : has_top (order_dual α) := ⟨(⊥ : α)⟩ -instance [has_top α] : has_bot (order_dual α) := ⟨(⊤ : α)⟩ +instance [has_bot α] : has_top αᵒᵈ := ⟨(⊥ : α)⟩ +instance [has_top α] : has_bot αᵒᵈ := ⟨(⊤ : α)⟩ -instance [has_le α] [order_bot α] : order_top (order_dual α) := +instance [has_le α] [order_bot α] : order_top αᵒᵈ := { le_top := @bot_le α _ _, .. order_dual.has_top α } -instance [has_le α] [order_top α] : order_bot (order_dual α) := +instance [has_le α] [order_top α] : order_bot αᵒᵈ := { bot_le := @le_top α _ _, .. order_dual.has_bot α } -instance [has_le α] [bounded_order α] : bounded_order (order_dual α) := +instance [has_le α] [bounded_order α] : bounded_order αᵒᵈ := { .. order_dual.order_top α, .. order_dual.order_bot α } end order_dual @@ -1221,7 +1221,7 @@ lemma max_top_right [order_top α] (a : α) : max a ⊤ = ⊤ := max_eq_right le by { symmetry, cases le_total a b; simpa [*, min_eq_left, min_eq_right] using eq_bot_mono h } @[simp] lemma max_eq_top [order_top α] {a b : α} : max a b = ⊤ ↔ a = ⊤ ∨ b = ⊤ := -@min_eq_bot (order_dual α) _ _ a b +@min_eq_bot αᵒᵈ _ _ a b @[simp] lemma max_eq_bot [order_bot α] {a b : α} : max a b = ⊥ ↔ a = ⊥ ∧ b = ⊥ := sup_eq_bot_iff @[simp] lemma min_eq_top [order_top α] {a b : α} : min a b = ⊤ ↔ a = ⊤ ∧ b = ⊤ := inf_eq_top_iff @@ -1408,7 +1408,7 @@ export is_complemented (exists_is_compl) namespace is_complemented variables [lattice α] [bounded_order α] [is_complemented α] -instance : is_complemented (order_dual α) := +instance : is_complemented αᵒᵈ := ⟨λ a, let ⟨b, hb⟩ := exists_is_compl (show α, from a) in ⟨b, hb.to_order_dual⟩⟩ end is_complemented diff --git a/src/order/bounds.lean b/src/order/bounds.lean index 8e432601bbf1a..d9972992fe65d 100644 --- a/src/order/bounds.lean +++ b/src/order/bounds.lean @@ -78,8 +78,7 @@ by simp [bdd_above, upper_bounds, set.nonempty] /-- A set `s` is not bounded below if and only if for each `x` there exists `y ∈ s` such that `x` is not less than or equal to `y`. This version only assumes `preorder` structure and uses `¬(x ≤ y)`. A version for linear orders is called `not_bdd_below_iff`. -/ -lemma not_bdd_below_iff' : ¬bdd_below s ↔ ∀ x, ∃ y ∈ s, ¬(x ≤ y) := -@not_bdd_above_iff' (order_dual α) _ _ +lemma not_bdd_below_iff' : ¬bdd_below s ↔ ∀ x, ∃ y ∈ s, ¬ x ≤ y := @not_bdd_above_iff' αᵒᵈ _ _ /-- A set `s` is not bounded above if and only if for each `x` there exists `y ∈ s` that is greater than `x`. A version for preorders is called `not_bdd_above_iff'`. -/ @@ -91,7 +90,7 @@ by simp only [not_bdd_above_iff', not_le] than `x`. A version for preorders is called `not_bdd_below_iff'`. -/ lemma not_bdd_below_iff {α : Type*} [linear_order α] {s : set α} : ¬bdd_below s ↔ ∀ x, ∃ y ∈ s, y < x := -@not_bdd_above_iff (order_dual α) _ _ +@not_bdd_above_iff αᵒᵈ _ _ lemma bdd_above.dual (h : bdd_above s) : bdd_below (of_dual ⁻¹' s) := h @@ -214,7 +213,7 @@ lemma is_lub_iff_le_iff : is_lub s a ↔ ∀ b, a ≤ b ↔ b ∈ upper_bounds s ⟨λ h b, is_lub_le_iff h, λ H, ⟨(H _).1 le_rfl, λ b hb, (H b).2 hb⟩⟩ lemma is_glb_iff_le_iff : is_glb s a ↔ ∀ b, b ≤ a ↔ b ∈ lower_bounds s := -@is_lub_iff_le_iff (order_dual α) _ _ _ +@is_lub_iff_le_iff αᵒᵈ _ _ _ /-- If `s` has a least upper bound, then it is bounded above. -/ lemma is_lub.bdd_above (h : is_lub s a) : bdd_above s := ⟨a, h.1⟩ @@ -242,7 +241,7 @@ subset.antisymm (λ b hb x hx, hx.elim (λ hs, hb.1 hs) (λ ht, hb.2 ht)) @[simp] lemma lower_bounds_union : lower_bounds (s ∪ t) = lower_bounds s ∩ lower_bounds t := -@upper_bounds_union (order_dual α) _ s t +@upper_bounds_union αᵒᵈ _ s t lemma union_upper_bounds_subset_upper_bounds_inter : upper_bounds s ∪ upper_bounds t ⊆ upper_bounds (s ∩ t) := @@ -252,7 +251,7 @@ union_subset lemma union_lower_bounds_subset_lower_bounds_inter : lower_bounds s ∪ lower_bounds t ⊆ lower_bounds (s ∩ t) := -@union_upper_bounds_subset_upper_bounds_inter (order_dual α) _ s t +@union_upper_bounds_subset_upper_bounds_inter αᵒᵈ _ s t lemma is_least_union_iff {a : α} {s t : set α} : is_least (s ∪ t) a ↔ (is_least s a ∧ a ∈ lower_bounds t ∨ a ∈ lower_bounds s ∧ is_least t a) := @@ -261,7 +260,7 @@ by simp [is_least, lower_bounds_union, or_and_distrib_right, and_comm (a ∈ t), lemma is_greatest_union_iff : is_greatest (s ∪ t) a ↔ (is_greatest s a ∧ a ∈ upper_bounds t ∨ a ∈ upper_bounds s ∧ is_greatest t a) := -@is_least_union_iff (order_dual α) _ a s t +@is_least_union_iff αᵒᵈ _ a s t /-- If `s` is bounded, then so is `s ∩ t` -/ lemma bdd_above.inter_of_left (h : bdd_above s) : bdd_above (s ∩ t) := @@ -298,12 +297,12 @@ lemma bdd_above_union [semilattice_sup γ] {s t : set γ} : lemma bdd_below.union [semilattice_inf γ] {s t : set γ} : bdd_below s → bdd_below t → bdd_below (s ∪ t) := -@bdd_above.union (order_dual γ) _ s t +@bdd_above.union γᵒᵈ _ s t /--The union of two sets is bounded above if and only if each of the sets is.-/ lemma bdd_below_union [semilattice_inf γ] {s t : set γ} : bdd_below (s ∪ t) ↔ bdd_below s ∧ bdd_below t := -@bdd_above_union (order_dual γ) _ s t +@bdd_above_union γᵒᵈ _ s t /-- If `a` is the least upper bound of `s` and `b` is the least upper bound of `t`, then `a ⊔ b` is the least upper bound of `s ∪ t`. -/ @@ -389,8 +388,7 @@ lemma bdd_below_Ioi : bdd_below (Ioi a) := ⟨a, λ x hx, le_of_lt hx⟩ lemma lub_Iio_le (a : α) (hb : is_lub (set.Iio a) b) : b ≤ a := (is_lub_le_iff hb).mpr $ λ k hk, le_of_lt hk -lemma le_glb_Ioi (a : α) (hb : is_glb (set.Ioi a) b) : a ≤ b := -@lub_Iio_le (order_dual α) _ _ a hb +lemma le_glb_Ioi (a : α) (hb : is_glb (set.Ioi a) b) : a ≤ b := @lub_Iio_le αᵒᵈ _ _ a hb lemma lub_Iio_eq_self_or_Iio_eq_Iic [partial_order γ] {j : γ} (i : γ) (hj : is_lub (set.Iio i) j) : j = i ∨ set.Iio i = set.Iic j := @@ -403,7 +401,7 @@ end lemma glb_Ioi_eq_self_or_Ioi_eq_Ici [partial_order γ] {j : γ} (i : γ) (hj : is_glb (set.Ioi i) j) : j = i ∨ set.Ioi i = set.Ici j := -@lub_Iio_eq_self_or_Iio_eq_Iic (order_dual γ) _ j i hj +@lub_Iio_eq_self_or_Iio_eq_Iic γᵒᵈ _ j i hj section @@ -422,15 +420,14 @@ begin exact h, }, end -lemma exists_glb_Ioi (i : γ) : ∃ j, is_glb (set.Ioi i) j := -@exists_lub_Iio (order_dual γ) _ i +lemma exists_glb_Ioi (i : γ) : ∃ j, is_glb (set.Ioi i) j := @exists_lub_Iio γᵒᵈ _ i variables [densely_ordered γ] lemma is_lub_Iio {a : γ} : is_lub (Iio a) a := ⟨λ x hx, le_of_lt hx, λ y hy, le_of_forall_ge_of_dense hy⟩ -lemma is_glb_Ioi {a : γ} : is_glb (Ioi a) a := @is_lub_Iio (order_dual γ) _ _ a +lemma is_glb_Ioi {a : γ} : is_glb (Ioi a) a := @is_lub_Iio γᵒᵈ _ _ a lemma upper_bounds_Iio {a : γ} : upper_bounds (Iio a) = Ici a := is_lub_Iio.upper_bounds_eq @@ -445,8 +442,7 @@ end lemma is_greatest_singleton : is_greatest {a} a := ⟨mem_singleton a, λ x hx, le_of_eq $ eq_of_mem_singleton hx⟩ -lemma is_least_singleton : is_least {a} a := -@is_greatest_singleton (order_dual α) _ a +lemma is_least_singleton : is_least {a} a := @is_greatest_singleton αᵒᵈ _ a lemma is_lub_singleton : is_lub {a} a := is_greatest_singleton.is_lub @@ -583,10 +579,10 @@ is_greatest_univ.is_lub @[simp] lemma order_bot.lower_bounds_univ [partial_order γ] [order_bot γ] : lower_bounds (univ : set γ) = {⊥} := -@order_top.upper_bounds_univ (order_dual γ) _ _ +@order_top.upper_bounds_univ γᵒᵈ _ _ lemma is_least_univ [preorder γ] [order_bot γ] : is_least (univ : set γ) ⊥ := -@is_greatest_univ (order_dual γ) _ _ +@is_greatest_univ γᵒᵈ _ _ lemma is_glb_univ [preorder γ] [order_bot γ] : is_glb (univ : set γ) ⊥ := is_least_univ.is_glb @@ -596,13 +592,13 @@ eq_empty_of_subset_empty $ λ b hb, let ⟨x, hx⟩ := exists_gt b in not_le_of_lt hx (hb trivial) @[simp] lemma no_min_order.lower_bounds_univ [no_min_order α] : lower_bounds (univ : set α) = ∅ := -@no_max_order.upper_bounds_univ (order_dual α) _ _ +@no_max_order.upper_bounds_univ αᵒᵈ _ _ @[simp] lemma not_bdd_above_univ [no_max_order α] : ¬bdd_above (univ : set α) := by simp [bdd_above] @[simp] lemma not_bdd_below_univ [no_min_order α] : ¬bdd_below (univ : set α) := -@not_bdd_above_univ (order_dual α) _ _ +@not_bdd_above_univ αᵒᵈ _ _ /-! #### Empty set @@ -611,8 +607,7 @@ by simp [bdd_above] @[simp] lemma upper_bounds_empty : upper_bounds (∅ : set α) = univ := by simp only [upper_bounds, eq_univ_iff_forall, mem_set_of_eq, ball_empty_iff, forall_true_iff] -@[simp] lemma lower_bounds_empty : lower_bounds (∅ : set α) = univ := -@upper_bounds_empty (order_dual α) _ +@[simp] lemma lower_bounds_empty : lower_bounds (∅ : set α) = univ := @upper_bounds_empty αᵒᵈ _ @[simp] lemma bdd_above_empty [nonempty α] : bdd_above (∅ : set α) := by simp only [bdd_above, upper_bounds_empty, univ_nonempty] @@ -623,8 +618,7 @@ by simp only [bdd_below, lower_bounds_empty, univ_nonempty] lemma is_glb_empty [preorder γ] [order_top γ] : is_glb ∅ (⊤:γ) := by simp only [is_glb, lower_bounds_empty, is_greatest_univ] -lemma is_lub_empty [preorder γ] [order_bot γ] : is_lub ∅ (⊥:γ) := -@is_glb_empty (order_dual γ) _ _ +lemma is_lub_empty [preorder γ] [order_bot γ] : is_lub ∅ (⊥:γ) := @is_glb_empty γᵒᵈ _ _ lemma is_lub.nonempty [no_min_order α] (hs : is_lub s a) : s.nonempty := let ⟨a', ha'⟩ := exists_lt a in @@ -638,7 +632,7 @@ lemma nonempty_of_not_bdd_above [ha : nonempty α] (h : ¬bdd_above s) : s.nonem nonempty.elim ha $ λ x, (not_bdd_above_iff'.1 h x).imp $ λ a ha, ha.fst lemma nonempty_of_not_bdd_below [ha : nonempty α] (h : ¬bdd_below s) : s.nonempty := -@nonempty_of_not_bdd_above (order_dual α) _ _ _ h +@nonempty_of_not_bdd_above αᵒᵈ _ _ _ h /-! #### insert @@ -718,7 +712,7 @@ is_greatest_singleton.insert _ ⟨λ H, ⟨λ x hx, H.2 $ subset_upper_bounds_lower_bounds s hx, H.1⟩, is_greatest.is_lub⟩ @[simp] lemma is_glb_upper_bounds : is_glb (upper_bounds s) a ↔ is_lub s a := -@is_lub_lower_bounds (order_dual α) _ _ _ +@is_lub_lower_bounds αᵒᵈ _ _ _ end @@ -1038,7 +1032,7 @@ lemma is_glb.of_image [preorder α] [preorder β] {f : α → β} (hf : ∀ {x y lemma is_lub.of_image [preorder α] [preorder β] {f : α → β} (hf : ∀ {x y}, f x ≤ f y ↔ x ≤ y) {s : set α} {x : α} (hx : is_lub (f '' s) (f x)) : is_lub s x := -@is_glb.of_image (order_dual α) (order_dual β) _ _ f (λ x y, hf) _ _ hx +@is_glb.of_image αᵒᵈ βᵒᵈ _ _ f (λ x y, hf) _ _ hx lemma is_lub_pi {π : α → Type*} [Π a, preorder (π a)] {s : set (Π a, π a)} {f : Π a, π a} : is_lub s f ↔ ∀ a, is_lub (function.eval a '' s) (f a) := @@ -1054,7 +1048,7 @@ end lemma is_glb_pi {π : α → Type*} [Π a, preorder (π a)] {s : set (Π a, π a)} {f : Π a, π a} : is_glb s f ↔ ∀ a, is_glb (function.eval a '' s) (f a) := -@is_lub_pi α (λ a, order_dual (π a)) _ s f +@is_lub_pi α (λ a, (π a)ᵒᵈ) _ s f lemma is_lub_prod [preorder α] [preorder β] {s : set (α × β)} (p : α × β) : is_lub s p ↔ is_lub (prod.fst '' s) p.1 ∧ is_lub (prod.snd '' s) p.2 := @@ -1072,7 +1066,7 @@ end lemma is_glb_prod [preorder α] [preorder β] {s : set (α × β)} (p : α × β) : is_glb s p ↔ is_glb (prod.fst '' s) p.1 ∧ is_glb (prod.snd '' s) p.2 := -@is_lub_prod (order_dual α) (order_dual β) _ _ _ _ +@is_lub_prod αᵒᵈ βᵒᵈ _ _ _ _ namespace order_iso @@ -1084,9 +1078,8 @@ subset.antisymm (λ x hx, ⟨f.symm x, λ y hy, f.le_symm_apply.2 (hx $ mem_image_of_mem _ hy), f.apply_symm_apply x⟩) f.monotone.image_upper_bounds_subset_upper_bounds_image -lemma lower_bounds_image {s : set α} : - lower_bounds (f '' s) = f '' lower_bounds s := -@upper_bounds_image (order_dual α) (order_dual β) _ _ f.dual _ +lemma lower_bounds_image {s : set α} : lower_bounds (f '' s) = f '' lower_bounds s := +@upper_bounds_image αᵒᵈ βᵒᵈ _ _ f.dual _ @[simp] lemma is_lub_image {s : set α} {x : β} : is_lub (f '' s) x ↔ is_lub s (f.symm x) := diff --git a/src/order/category/BoolAlg.lean b/src/order/category/BoolAlg.lean index 13894871ed017..7301e685995a9 100644 --- a/src/order/category/BoolAlg.lean +++ b/src/order/category/BoolAlg.lean @@ -53,7 +53,7 @@ between them. -/ /-- `order_dual` as a functor. -/ @[simps] def dual : BoolAlg ⥤ BoolAlg := -{ obj := λ X, of (order_dual X), map := λ X Y, bounded_lattice_hom.dual } +{ obj := λ X, of Xᵒᵈ, map := λ X Y, bounded_lattice_hom.dual } /-- The equivalence between `BoolAlg` and itself induced by `order_dual` both ways. -/ @[simps functor inverse] def dual_equiv : BoolAlg ≌ BoolAlg := diff --git a/src/order/category/BoundedDistribLattice.lean b/src/order/category/BoundedDistribLattice.lean index 58e342ecbddf9..ecac2f1ba4af5 100644 --- a/src/order/category/BoundedDistribLattice.lean +++ b/src/order/category/BoundedDistribLattice.lean @@ -68,7 +68,7 @@ between them. -/ /-- `order_dual` as a functor. -/ @[simps] def dual : BoundedDistribLattice ⥤ BoundedDistribLattice := -{ obj := λ X, of (order_dual X), map := λ X Y, bounded_lattice_hom.dual } +{ obj := λ X, of Xᵒᵈ, map := λ X Y, bounded_lattice_hom.dual } /-- The equivalence between `BoundedDistribLattice` and itself induced by `order_dual` both ways. -/ @[simps functor inverse] def dual_equiv : BoundedDistribLattice ≌ BoundedDistribLattice := diff --git a/src/order/category/BoundedLattice.lean b/src/order/category/BoundedLattice.lean index 01043558a12a1..0201a28d1a191 100644 --- a/src/order/category/BoundedLattice.lean +++ b/src/order/category/BoundedLattice.lean @@ -98,7 +98,7 @@ between them. -/ /-- `order_dual` as a functor. -/ @[simps] def dual : BoundedLattice ⥤ BoundedLattice := -{ obj := λ X, of (order_dual X), map := λ X Y, bounded_lattice_hom.dual } +{ obj := λ X, of Xᵒᵈ, map := λ X Y, bounded_lattice_hom.dual } /-- The equivalence between `BoundedLattice` and itself induced by `order_dual` both ways. -/ @[simps functor inverse] def dual_equiv : BoundedLattice ≌ BoundedLattice := diff --git a/src/order/category/BoundedOrder.lean b/src/order/category/BoundedOrder.lean index 5db92d0feda2b..8c359e2282b31 100644 --- a/src/order/category/BoundedOrder.lean +++ b/src/order/category/BoundedOrder.lean @@ -57,7 +57,7 @@ instance has_forget_to_Bipointed : has_forget₂ BoundedOrder Bipointed := /-- `order_dual` as a functor. -/ @[simps] def dual : BoundedOrder ⥤ BoundedOrder := -{ obj := λ X, of (order_dual X), map := λ X Y, bounded_order_hom.dual } +{ obj := λ X, of Xᵒᵈ, map := λ X Y, bounded_order_hom.dual } /-- Constructs an equivalence between bounded orders from an order isomorphism between them. -/ @[simps] def iso.mk {α β : BoundedOrder.{u}} (e : α ≃o β) : α ≅ β := diff --git a/src/order/category/CompleteLattice.lean b/src/order/category/CompleteLattice.lean index ffc4930c60e76..03ce1f1839b3a 100644 --- a/src/order/category/CompleteLattice.lean +++ b/src/order/category/CompleteLattice.lean @@ -53,7 +53,7 @@ instance has_forget_to_BoundedLattice : has_forget₂ CompleteLattice BoundedLat /-- `order_dual` as a functor. -/ @[simps] def dual : CompleteLattice ⥤ CompleteLattice := -{ obj := λ X, of (order_dual X), map := λ X Y, complete_lattice_hom.dual } +{ obj := λ X, of Xᵒᵈ, map := λ X Y, complete_lattice_hom.dual } /-- The equivalence between `CompleteLattice` and itself induced by `order_dual` both ways. -/ @[simps functor inverse] def dual_equiv : CompleteLattice ≌ CompleteLattice := diff --git a/src/order/category/DistribLattice.lean b/src/order/category/DistribLattice.lean index b966c9f4c3099..7b7220477b2b9 100644 --- a/src/order/category/DistribLattice.lean +++ b/src/order/category/DistribLattice.lean @@ -50,7 +50,7 @@ instance has_forget_to_Lattice : has_forget₂ DistribLattice Lattice := bundled /-- `order_dual` as a functor. -/ @[simps] def dual : DistribLattice ⥤ DistribLattice := -{ obj := λ X, of (order_dual X), map := λ X Y, lattice_hom.dual } +{ obj := λ X, of Xᵒᵈ, map := λ X Y, lattice_hom.dual } /-- The equivalence between `DistribLattice` and itself induced by `order_dual` both ways. -/ @[simps functor inverse] def dual_equiv : DistribLattice ≌ DistribLattice := diff --git a/src/order/category/FinBoolAlg.lean b/src/order/category/FinBoolAlg.lean index 26981b8b6003d..a6a1960a6ed5d 100644 --- a/src/order/category/FinBoolAlg.lean +++ b/src/order/category/FinBoolAlg.lean @@ -76,7 +76,7 @@ them. -/ /-- `order_dual` as a functor. -/ @[simps] def dual : FinBoolAlg ⥤ FinBoolAlg := -{ obj := λ X, of (order_dual X), map := λ X Y, bounded_lattice_hom.dual } +{ obj := λ X, of Xᵒᵈ, map := λ X Y, bounded_lattice_hom.dual } /-- The equivalence between `FinBoolAlg` and itself induced by `order_dual` both ways. -/ @[simps functor inverse] def dual_equiv : FinBoolAlg ≌ FinBoolAlg := diff --git a/src/order/category/FinPartialOrder.lean b/src/order/category/FinPartialOrder.lean index 13186942f9072..78b7eae8b2fd9 100644 --- a/src/order/category/FinPartialOrder.lean +++ b/src/order/category/FinPartialOrder.lean @@ -64,7 +64,7 @@ instance has_forget_to_Fintype : has_forget₂ FinPartialOrder Fintype := /-- `order_dual` as a functor. -/ @[simps] def dual : FinPartialOrder ⥤ FinPartialOrder := -{ obj := λ X, of (order_dual X), map := λ X Y, order_hom.dual } +{ obj := λ X, of Xᵒᵈ, map := λ X Y, order_hom.dual } /-- The equivalence between `FinPartialOrder` and itself induced by `order_dual` both ways. -/ @[simps functor inverse] def dual_equiv : FinPartialOrder ≌ FinPartialOrder := diff --git a/src/order/category/Lattice.lean b/src/order/category/Lattice.lean index ff1641b860e3d..3fa08568307f4 100644 --- a/src/order/category/Lattice.lean +++ b/src/order/category/Lattice.lean @@ -60,8 +60,7 @@ instance has_forget_to_PartialOrder : has_forget₂ Lattice PartialOrder := inv_hom_id' := by { ext, exact e.apply_symm_apply _ } } /-- `order_dual` as a functor. -/ -@[simps] def dual : Lattice ⥤ Lattice := -{ obj := λ X, of (order_dual X), map := λ X Y, lattice_hom.dual } +@[simps] def dual : Lattice ⥤ Lattice := { obj := λ X, of Xᵒᵈ, map := λ X Y, lattice_hom.dual } /-- The equivalence between `Lattice` and itself induced by `order_dual` both ways. -/ @[simps functor inverse] def dual_equiv : Lattice ≌ Lattice := diff --git a/src/order/category/LinearOrder.lean b/src/order/category/LinearOrder.lean index 97107e1468b93..e8d0bda96df25 100644 --- a/src/order/category/LinearOrder.lean +++ b/src/order/category/LinearOrder.lean @@ -49,7 +49,7 @@ instance has_forget_to_Lattice : has_forget₂ LinearOrder Lattice := /-- `order_dual` as a functor. -/ @[simps] def dual : LinearOrder ⥤ LinearOrder := -{ obj := λ X, of (order_dual X), map := λ X Y, order_hom.dual } +{ obj := λ X, of Xᵒᵈ, map := λ X Y, order_hom.dual } /-- The equivalence between `LinearOrder` and itself induced by `order_dual` both ways. -/ @[simps functor inverse] def dual_equiv : LinearOrder ≌ LinearOrder := diff --git a/src/order/category/NonemptyFinLinOrd.lean b/src/order/category/NonemptyFinLinOrd.lean index 01d1e8e7b8b32..d281dff7e3125 100644 --- a/src/order/category/NonemptyFinLinOrd.lean +++ b/src/order/category/NonemptyFinLinOrd.lean @@ -42,7 +42,7 @@ instance ulift.nonempty_fin_lin_ord (α : Type u) [nonempty_fin_lin_ord α] : .. linear_order.lift equiv.ulift (equiv.injective _), .. ulift.fintype _ } -instance (α : Type*) [nonempty_fin_lin_ord α] : nonempty_fin_lin_ord (order_dual α) := +instance (α : Type*) [nonempty_fin_lin_ord α] : nonempty_fin_lin_ord αᵒᵈ := { ..order_dual.fintype α } /-- The category of nonempty finite linear orders. -/ @@ -78,7 +78,7 @@ between them. -/ /-- `order_dual` as a functor. -/ @[simps] def dual : NonemptyFinLinOrd ⥤ NonemptyFinLinOrd := -{ obj := λ X, of (order_dual X), map := λ X Y, order_hom.dual } +{ obj := λ X, of Xᵒᵈ, map := λ X Y, order_hom.dual } /-- The equivalence between `FinPartialOrder` and itself induced by `order_dual` both ways. -/ @[simps functor inverse] def dual_equiv : NonemptyFinLinOrd ≌ NonemptyFinLinOrd := diff --git a/src/order/category/PartialOrder.lean b/src/order/category/PartialOrder.lean index 601fef3ed4de4..793f355aa681f 100644 --- a/src/order/category/PartialOrder.lean +++ b/src/order/category/PartialOrder.lean @@ -47,7 +47,7 @@ instance has_forget_to_Preorder : has_forget₂ PartialOrder Preorder := bundled /-- `order_dual` as a functor. -/ @[simps] def dual : PartialOrder ⥤ PartialOrder := -{ obj := λ X, of (order_dual X), map := λ X Y, order_hom.dual } +{ obj := λ X, of Xᵒᵈ, map := λ X Y, order_hom.dual } /-- The equivalence between `PartialOrder` and itself induced by `order_dual` both ways. -/ @[simps functor inverse] def dual_equiv : PartialOrder ≌ PartialOrder := diff --git a/src/order/category/Preorder.lean b/src/order/category/Preorder.lean index 5517cfe0644f8..b28dd988e00b2 100644 --- a/src/order/category/Preorder.lean +++ b/src/order/category/Preorder.lean @@ -52,7 +52,7 @@ instance (α : Preorder) : preorder α := α.str /-- `order_dual` as a functor. -/ @[simps] def dual : Preorder ⥤ Preorder := -{ obj := λ X, of (order_dual X), map := λ X Y, order_hom.dual } +{ obj := λ X, of Xᵒᵈ, map := λ X Y, order_hom.dual } /-- The equivalence between `Preorder` and itself induced by `order_dual` both ways. -/ @[simps functor inverse] def dual_equiv : Preorder ≌ Preorder := diff --git a/src/order/category/Semilattice.lean b/src/order/category/Semilattice.lean index 743f255f090eb..234a24b254b47 100644 --- a/src/order/category/Semilattice.lean +++ b/src/order/category/Semilattice.lean @@ -112,7 +112,7 @@ namespace SemilatticeSup /-- `order_dual` as a functor. -/ @[simps] def dual : SemilatticeSup ⥤ SemilatticeInf := -{ obj := λ X, SemilatticeInf.of (order_dual X), map := λ X Y, sup_bot_hom.dual } +{ obj := λ X, SemilatticeInf.of Xᵒᵈ, map := λ X Y, sup_bot_hom.dual } end SemilatticeSup @@ -127,7 +127,7 @@ namespace SemilatticeInf /-- `order_dual` as a functor. -/ @[simps] def dual : SemilatticeInf ⥤ SemilatticeSup := -{ obj := λ X, SemilatticeSup.of (order_dual X), map := λ X Y, inf_top_hom.dual } +{ obj := λ X, SemilatticeSup.of Xᵒᵈ, map := λ X Y, inf_top_hom.dual } end SemilatticeInf diff --git a/src/order/circular.lean b/src/order/circular.lean index 823316bc3056b..5255e3f546214 100644 --- a/src/order/circular.lean +++ b/src/order/circular.lean @@ -48,12 +48,12 @@ Some concrete circular orders one encounters in the wild are `zmod n` for `0 < n ## Notes -There's an unsolved diamond here. The instances `has_le α → has_btw (order_dual α)` and -`has_lt α → has_sbtw (order_dual α)` can each be inferred in two ways: -* `has_le α` → `has_btw α` → `has_btw (order_dual α)` vs - `has_le α` → `has_le (order_dual α)` → `has_btw (order_dual α)` -* `has_lt α` → `has_sbtw α` → `has_sbtw (order_dual α)` vs - `has_lt α` → `has_lt (order_dual α)` → `has_sbtw (order_dual α)` +There's an unsolved diamond on `order_dual α` here. The instances `has_le α → has_btw αᵒᵈ` and +`has_lt α → has_sbtw αᵒᵈ` can each be inferred in two ways: +* `has_le α` → `has_btw α` → `has_btw αᵒᵈ` vs + `has_le α` → `has_le αᵒᵈ` → `has_btw αᵒᵈ` +* `has_lt α` → `has_sbtw α` → `has_sbtw αᵒᵈ` vs + `has_lt α` → `has_lt αᵒᵈ` → `has_sbtw αᵒᵈ` The fields are propeq, but not defeq. It is temporarily fixed by turning the circularizing instances into definitions. @@ -385,10 +385,10 @@ def linear_order.to_circular_order (α : Type*) [linear_order α] : section order_dual -instance (α : Type*) [has_btw α] : has_btw (order_dual α) := ⟨λ a b c : α, btw c b a⟩ -instance (α : Type*) [has_sbtw α] : has_sbtw (order_dual α) := ⟨λ a b c : α, sbtw c b a⟩ +instance (α : Type*) [has_btw α] : has_btw αᵒᵈ := ⟨λ a b c : α, btw c b a⟩ +instance (α : Type*) [has_sbtw α] : has_sbtw αᵒᵈ := ⟨λ a b c : α, sbtw c b a⟩ -instance (α : Type*) [h : circular_preorder α] : circular_preorder (order_dual α) := +instance (α : Type*) [h : circular_preorder α] : circular_preorder αᵒᵈ := { btw_refl := btw_refl, btw_cyclic_left := λ a b c, btw_cyclic_right, sbtw_trans_left := λ a b c d habc hbdc, hbdc.trans_right habc, @@ -396,11 +396,11 @@ instance (α : Type*) [h : circular_preorder α] : circular_preorder (order_dual .. order_dual.has_btw α, .. order_dual.has_sbtw α } -instance (α : Type*) [circular_partial_order α] : circular_partial_order (order_dual α) := +instance (α : Type*) [circular_partial_order α] : circular_partial_order αᵒᵈ := { btw_antisymm := λ a b c habc hcba, @btw_antisymm α _ _ _ _ hcba habc, .. order_dual.circular_preorder α } -instance (α : Type*) [circular_order α] : circular_order (order_dual α) := +instance (α : Type*) [circular_order α] : circular_order αᵒᵈ := { btw_total := λ a b c, btw_total c b a, .. order_dual.circular_partial_order α } end order_dual diff --git a/src/order/compare.lean b/src/order/compare.lean index b792a468854a3..a9cefd8331663 100644 --- a/src/order/compare.lean +++ b/src/order/compare.lean @@ -127,7 +127,7 @@ by cases o₁; cases o₂; exact dec_trivial end ordering lemma order_dual.dual_compares [has_lt α] {a b : α} {o : ordering} : - @ordering.compares (order_dual α) _ o a b ↔ @ordering.compares α _ o b a := + @ordering.compares αᵒᵈ _ o a b ↔ @ordering.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 := @@ -146,7 +146,7 @@ begin end lemma order_dual.cmp_le_flip {α} [has_le α] [@decidable_rel α (≤)] (x y : α) : - @cmp_le (order_dual α) _ _ x y = cmp_le y x := rfl + @cmp_le αᵒᵈ _ _ x y = cmp_le y x := rfl /-- Generate a linear order structure from a preorder and `cmp` function. -/ def linear_order_of_compares [preorder α] (cmp : α → α → ordering) diff --git a/src/order/complete_boolean_algebra.lean b/src/order/complete_boolean_algebra.lean index 5270899009733..d084d45c86a3a 100644 --- a/src/order/complete_boolean_algebra.lean +++ b/src/order/complete_boolean_algebra.lean @@ -64,7 +64,7 @@ instance complete_distrib_lattice.to_coframe [complete_distrib_lattice α] : cof section frame variables [frame α] {s t : set α} {a b : α} -instance order_dual.coframe : coframe (order_dual α) := +instance order_dual.coframe : coframe αᵒᵈ := { infi_sup_le_sup_Inf := frame.inf_Sup_le_supr_inf, ..order_dual.complete_lattice α } lemma inf_Sup_eq : a ⊓ Sup s = ⨆ b ∈ s, a ⊓ b := @@ -122,37 +122,31 @@ end frame section coframe variables [coframe α] {s t : set α} {a b : α} -instance order_dual.frame : frame (order_dual α) := +instance order_dual.frame : frame αᵒᵈ := { inf_Sup_le_supr_inf := coframe.infi_sup_le_sup_Inf, ..order_dual.complete_lattice α } -theorem sup_Inf_eq : a ⊔ Inf s = (⨅ b ∈ s, a ⊔ b) := -@inf_Sup_eq (order_dual α) _ _ _ +lemma sup_Inf_eq : a ⊔ Inf s = ⨅ b ∈ s, a ⊔ b := @inf_Sup_eq αᵒᵈ _ _ _ +lemma Inf_sup_eq : Inf s ⊔ b = ⨅ a ∈ s, a ⊔ b := @Sup_inf_eq αᵒᵈ _ _ _ -theorem Inf_sup_eq : Inf s ⊔ b = (⨅ a ∈ s, a ⊔ b) := -@Sup_inf_eq (order_dual α) _ _ _ - -theorem infi_sup_eq (f : ι → α) (a : α) : (⨅ i, f i) ⊔ a = ⨅ i, f i ⊔ a := -@supr_inf_eq (order_dual α) _ _ _ _ - -theorem sup_infi_eq (a : α) (f : ι → α) : a ⊔ (⨅ i, f i) = ⨅ i, a ⊔ f i := -@inf_supr_eq (order_dual α) _ _ _ _ +lemma infi_sup_eq (f : ι → α) (a : α) : (⨅ i, f i) ⊔ a = ⨅ i, f i ⊔ a := @supr_inf_eq αᵒᵈ _ _ _ _ +lemma sup_infi_eq (a : α) (f : ι → α) : a ⊔ (⨅ i, f i) = ⨅ i, a ⊔ f i := @inf_supr_eq αᵒᵈ _ _ _ _ lemma binfi_sup_eq {f : Π i, κ i → α} (a : α) : (⨅ i j, f i j) ⊔ a = ⨅ i j, f i j ⊔ a := -@bsupr_inf_eq (order_dual α) _ _ _ _ _ +@bsupr_inf_eq αᵒᵈ _ _ _ _ _ lemma sup_binfi_eq {f : Π i, κ i → α} (a : α) : a ⊔ (⨅ i j, f i j) = ⨅ i j, a ⊔ f i j := -@inf_bsupr_eq (order_dual α) _ _ _ _ _ +@inf_bsupr_eq αᵒᵈ _ _ _ _ _ lemma infi_sup_infi {ι ι' : Type*} {f : ι → α} {g : ι' → α} : (⨅ i, f i) ⊔ (⨅ i, g i) = ⨅ i : ι × ι', f i.1 ⊔ g i.2 := -@supr_inf_supr (order_dual α) _ _ _ _ _ +@supr_inf_supr αᵒᵈ _ _ _ _ _ lemma binfi_sup_binfi {ι ι' : Type*} {f : ι → α} {g : ι' → α} {s : set ι} {t : set ι'} : (⨅ i ∈ s, f i) ⊔ (⨅ j ∈ t, g j) = ⨅ p ∈ s ×ˢ t, f (p : ι × ι').1 ⊔ g p.2 := -@bsupr_inf_bsupr (order_dual α) _ _ _ _ _ _ _ +@bsupr_inf_bsupr αᵒᵈ _ _ _ _ _ _ _ theorem Inf_sup_Inf : Inf s ⊔ Inf t = (⨅ p ∈ s ×ˢ t, (p : α × α).1 ⊔ p.2) := -@Sup_inf_Sup (order_dual α) _ _ _ +@Sup_inf_Sup αᵒᵈ _ _ _ instance pi.coframe {ι : Type*} {π : ι → Type*} [Π i, coframe (π i)] : coframe (Π i, π i) := { Inf := Inf, @@ -165,7 +159,7 @@ end coframe section complete_distrib_lattice variables [complete_distrib_lattice α] {a b : α} {s t : set α} -instance : complete_distrib_lattice (order_dual α) := { ..order_dual.frame, ..order_dual.coframe } +instance : complete_distrib_lattice αᵒᵈ := { ..order_dual.frame, ..order_dual.coframe } instance pi.complete_distrib_lattice {ι : Type*} {π : ι → Type*} [Π i, complete_distrib_lattice (π i)] : complete_distrib_lattice (Π i, π i) := diff --git a/src/order/complete_lattice.lean b/src/order/complete_lattice.lean index c299d80d2ebb6..e67a577f6f81f 100644 --- a/src/order/complete_lattice.lean +++ b/src/order/complete_lattice.lean @@ -67,8 +67,8 @@ def infi [has_Inf α] {ι} (s : ι → α) : α := Inf (range s) notation `⨆` binders `, ` r:(scoped f, supr f) := r notation `⨅` binders `, ` r:(scoped f, infi f) := r -instance (α) [has_Inf α] : has_Sup (order_dual α) := ⟨(Inf : set α → α)⟩ -instance (α) [has_Sup α] : has_Inf (order_dual α) := ⟨(Sup : set α → α)⟩ +instance (α) [has_Inf α] : has_Sup αᵒᵈ := ⟨(Inf : set α → α)⟩ +instance (α) [has_Sup α] : has_Inf αᵒᵈ := ⟨(Sup : set α → α)⟩ /-- Note that we rarely use `complete_semilattice_Sup` @@ -280,7 +280,7 @@ class complete_linear_order (α : Type*) extends complete_lattice α, namespace order_dual variable (α) -instance [complete_lattice α] : complete_lattice (order_dual α) := +instance [complete_lattice α] : complete_lattice αᵒᵈ := { le_Sup := @complete_lattice.Inf_le α _, Sup_le := @complete_lattice.le_Inf α _, Inf_le := @complete_lattice.le_Sup α _, @@ -288,7 +288,7 @@ instance [complete_lattice α] : complete_lattice (order_dual α) := .. order_dual.lattice α, ..order_dual.has_Sup α, ..order_dual.has_Inf α, .. order_dual.bounded_order α } -instance [complete_linear_order α] : complete_linear_order (order_dual α) := +instance [complete_linear_order α] : complete_linear_order αᵒᵈ := { .. order_dual.complete_lattice α, .. order_dual.linear_order α } end order_dual @@ -311,8 +311,7 @@ Sup_le $ λ b hb, le_inf (le_Sup hb.1) (le_Sup hb.2) theorem Inf_union {s t : set α} : Inf (s ∪ t) = Inf s ⊓ Inf t := ((is_glb_Inf s).union (is_glb_Inf t)).Inf_eq -theorem le_Inf_inter {s t : set α} : Inf s ⊔ Inf t ≤ Inf (s ∩ t) := -@Sup_inter_le (order_dual α) _ _ _ +theorem le_Inf_inter {s t : set α} : Inf s ⊔ Inf t ≤ Inf (s ∩ t) := @Sup_inter_le αᵒᵈ _ _ _ @[simp] theorem Sup_empty : Sup ∅ = (⊥ : α) := (@is_lub_empty α _ _).Sup_eq @@ -354,8 +353,7 @@ lemma eq_singleton_top_of_Inf_eq_top_of_nonempty {s : set α} (h_inf : Inf s = ⊤) (hne : s.nonempty) : s = {⊤} := by { rw set.eq_singleton_iff_nonempty_unique_mem, rw Inf_eq_top at h_inf, exact ⟨hne, h_inf⟩, } -@[simp] theorem Sup_eq_bot : Sup s = ⊥ ↔ (∀ a ∈ s, a = ⊥) := -@Inf_eq_top (order_dual α) _ _ +@[simp] lemma Sup_eq_bot : Sup s = ⊥ ↔ ∀ a ∈ s, a = ⊥ := @Inf_eq_top αᵒᵈ _ _ lemma eq_singleton_bot_of_Sup_eq_bot_of_nonempty {s : set α} (h_sup : Sup s = ⊥) (hne : s.nonempty) : s = {⊥} := @@ -380,7 +378,7 @@ See `cInf_eq_of_forall_ge_of_forall_gt_exists_lt` for a version in conditionally lattices. -/ theorem Inf_eq_of_forall_ge_of_forall_gt_exists_lt : (∀ a ∈ s, b ≤ a) → (∀ w, b < w → (∃ a ∈ s, a < w)) → Inf s = b := -@Sup_eq_of_forall_le_of_forall_lt_exists_gt (order_dual α) _ _ ‹_› +@Sup_eq_of_forall_le_of_forall_lt_exists_gt αᵒᵈ _ _ ‹_› end @@ -397,8 +395,7 @@ iff.intro let ⟨a, ha, h⟩ := h _ h' in lt_irrefl a $ lt_of_le_of_lt (le_Sup ha) h) -lemma Inf_eq_bot : Inf s = ⊥ ↔ (∀ b > ⊥, ∃ a ∈ s, a < b) := -@Sup_eq_top (order_dual α) _ _ +lemma Inf_eq_bot : Inf s = ⊥ ↔ ∀ b > ⊥, ∃ a ∈ s, a < b := @Sup_eq_top αᵒᵈ _ _ lemma lt_supr_iff {f : ι → α} : a < supr f ↔ ∃ i, a < f i := lt_Sup_iff.trans exists_range_iff lemma infi_lt_iff {f : ι → α} : infi f < a ↔ ∃ i, f i < a := Inf_lt_iff.trans exists_range_iff @@ -445,27 +442,26 @@ section has_Inf variables [has_Inf α] {f g : ι → α} lemma Inf_range : Inf (range f) = infi f := rfl -lemma Inf_eq_infi' (s : set α) : Inf s = ⨅ a : s, a := @Sup_eq_supr' (order_dual α) _ _ +lemma Inf_eq_infi' (s : set α) : Inf s = ⨅ a : s, a := @Sup_eq_supr' αᵒᵈ _ _ lemma infi_congr (h : ∀ i, f i = g i) : (⨅ i, f i) = ⨅ i, g i := congr_arg _ $ funext h lemma function.surjective.infi_comp {f : ι → ι'} (hf : surjective f) (g : ι' → α) : (⨅ x, g (f x)) = ⨅ y, g y := -@function.surjective.supr_comp (order_dual α) _ _ _ f hf g +@function.surjective.supr_comp αᵒᵈ _ _ _ f hf g lemma function.surjective.infi_congr {g : ι' → α} (h : ι → ι') (h1 : surjective h) (h2 : ∀ x, g (h x) = f x) : (⨅ x, f x) = ⨅ y, g y := -@function.surjective.supr_congr (order_dual α) _ _ _ _ _ h h1 h2 +@function.surjective.supr_congr αᵒᵈ _ _ _ _ _ h h1 h2 @[congr]lemma infi_congr_Prop {p q : Prop} {f₁ : p → α} {f₂ : q → α} (pq : p ↔ q) (f : ∀ x, f₁ (pq.mpr x) = f₂ x) : infi f₁ = infi f₂ := -@supr_congr_Prop (order_dual α) _ p q f₁ f₂ pq f +@supr_congr_Prop αᵒᵈ _ p q f₁ f₂ pq f lemma infi_range' (g : β → α) (f : ι → β) : (⨅ b : range f, g b) = ⨅ i, g (f i) := -@supr_range' (order_dual α) _ _ _ _ _ +@supr_range' αᵒᵈ _ _ _ _ _ -lemma Inf_image' {s : set β} {f : β → α} : Inf (f '' s) = (⨅ a : s, f a) := -@Sup_image' (order_dual α) _ _ _ _ +lemma Inf_image' {s : set β} {f : β → α} : Inf (f '' s) = ⨅ a : s, f a := @Sup_image' αᵒᵈ _ _ _ _ end has_Inf @@ -577,8 +573,7 @@ lemma lt_infi_iff : a < infi f ↔ ∃ b, a < b ∧ ∀ i, b ≤ f i := lemma Sup_eq_supr {s : set α} : Sup s = ⨆ a ∈ s, a := le_antisymm (Sup_le le_supr₂) (supr₂_le $ λ b, le_Sup) -lemma Inf_eq_infi {s : set α} : Inf s = ⨅ a ∈ s, a := -@Sup_eq_supr (order_dual α) _ _ +lemma Inf_eq_infi {s : set α} : Inf s = ⨅ a ∈ s, a := @Sup_eq_supr αᵒᵈ _ _ lemma Sup_sUnion (s : set (set α)) : Sup (⋃₀ s) = ⨆ t ∈ s, Sup t := begin @@ -591,7 +586,7 @@ begin exact supr_le (λ ts, Sup_le_Sup (λ x xt, ⟨t, ts, xt⟩)) } end -lemma Inf_sUnion (s : set (set α)) : Inf (⋃₀ s) = ⨅ t ∈ s, Inf t := @Sup_sUnion (order_dual α) _ _ +lemma Inf_sUnion (s : set (set α)) : Inf (⋃₀ s) = ⨅ t ∈ s, Inf t := @Sup_sUnion αᵒᵈ _ _ lemma monotone.le_map_supr [complete_lattice β] {f : α → β} (hf : monotone f) : (⨆ i, f (s i)) ≤ f (supr s) := @@ -685,14 +680,10 @@ theorem infi_const [nonempty ι] {a : α} : (⨅ b : ι, a) = a := by rw [infi, range_const, Inf_singleton] -- We will generalize this to conditionally complete lattices in `csupr_const`. -theorem supr_const [nonempty ι] {a : α} : (⨆ b : ι, a) = a := -@infi_const (order_dual α) _ _ _ _ +theorem supr_const [nonempty ι] {a : α} : (⨆ b : ι, a) = a := @infi_const αᵒᵈ _ _ _ _ -@[simp] lemma infi_top : (⨅ i:ι, ⊤ : α) = ⊤ := -top_unique $ le_infi $ λ i, le_rfl - -@[simp] lemma supr_bot : (⨆ i:ι, ⊥ : α) = ⊥ := -@infi_top (order_dual α) _ _ +@[simp] lemma infi_top : (⨅ i:ι, ⊤ : α) = ⊤ := top_unique $ le_infi $ λ i, le_rfl +@[simp] lemma supr_bot : (⨆ i:ι, ⊥ : α) = ⊥ := @infi_top αᵒᵈ _ _ @[simp] lemma supr_eq_bot : supr s = ⊥ ↔ ∀ i, s i = ⊥ := Sup_eq_bot.trans forall_range_iff @[simp] lemma infi_eq_top : infi s = ⊤ ↔ ∀ i, s i = ⊤ := Inf_eq_top.trans forall_range_iff @@ -730,7 +721,7 @@ See `cinfi_eq_of_forall_ge_of_forall_gt_exists_lt` for a version in conditionall lattices. -/ theorem infi_eq_of_forall_ge_of_forall_gt_exists_lt {f : ι → α} (h₁ : ∀ i, b ≤ f i) (h₂ : ∀ w, b < w → (∃ i, f i < w)) : (⨅ (i : ι), f i) = b := -@supr_eq_of_forall_le_of_forall_lt_exists_gt (order_dual α) _ _ _ ‹_› ‹_› ‹_› +@supr_eq_of_forall_le_of_forall_lt_exists_gt αᵒᵈ _ _ _ ‹_› ‹_› ‹_› lemma supr_eq_dif {p : Prop} [decidable p] (a : p → α) : (⨆ h : p, a h) = if h : p then a h else ⊥ := @@ -742,7 +733,7 @@ supr_eq_dif (λ _, a) lemma infi_eq_dif {p : Prop} [decidable p] (a : p → α) : (⨅ h : p, a h) = if h : p then a h else ⊤ := -@supr_eq_dif (order_dual α) _ _ _ _ +@supr_eq_dif αᵒᵈ _ _ _ _ lemma infi_eq_if {p : Prop} [decidable p] (a : α) : (⨅ h : p, a) = if p then a else ⊤ := @@ -755,8 +746,7 @@ le_antisymm (supr_le $ λ j, supr_mono $ λ i, le_supr _ _) -- TODO: should this be @[simp]? -lemma infi_comm {f : ι → ι' → α} : (⨅ i j, f i j) = ⨅ j i, f i j := -@supr_comm (order_dual α) _ _ _ _ +lemma infi_comm {f : ι → ι' → α} : (⨅ i j, f i j) = ⨅ j i, f i j := @supr_comm αᵒᵈ _ _ _ _ /- TODO: this is strange. In the proof below, we get exactly the desired among the equalities, but close does not get it. @@ -784,11 +774,11 @@ le_antisymm @[simp] theorem supr_supr_eq_left {b : β} {f : Π x : β, x = b → α} : (⨆ x, ⨆ h : x = b, f x h) = f b rfl := -@infi_infi_eq_left (order_dual α) _ _ _ _ +@infi_infi_eq_left αᵒᵈ _ _ _ _ @[simp] theorem supr_supr_eq_right {b : β} {f : Π x : β, b = x → α} : (⨆ x, ⨆ h : b = x, f x h) = f b rfl := -@infi_infi_eq_right (order_dual α) _ _ _ _ +@infi_infi_eq_right αᵒᵈ _ _ _ _ attribute [ematch] le_refl @@ -833,13 +823,13 @@ lemma inf_binfi {p : ι → Prop} {f : Π i (hi : p i), α} {a : α} (h : ∃ i, by simpa only [inf_comm] using binfi_inf h theorem supr_sup_eq {f g : ι → α} : (⨆ x, f x ⊔ g x) = (⨆ x, f x) ⊔ (⨆ x, g x) := -@infi_inf_eq (order_dual α) ι _ _ _ +@infi_inf_eq αᵒᵈ ι _ _ _ lemma supr_sup [h : nonempty ι] {f : ι → α} {a : α} : (⨆ x, f x) ⊔ a = (⨆ x, f x ⊔ a) := -@infi_inf (order_dual α) _ _ _ _ _ +@infi_inf αᵒᵈ _ _ _ _ _ lemma sup_supr [nonempty ι] {f : ι → α} {a : α} : a ⊔ (⨆ x, f x) = (⨆ x, a ⊔ f x) := -@inf_infi (order_dual α) _ _ _ _ _ +@inf_infi αᵒᵈ _ _ _ _ _ /-! ### `supr` and `infi` under `Prop` -/ @@ -856,7 +846,7 @@ lemma infi_true {s : true → α} : infi s = s trivial := infi_pos trivial le_antisymm (le_infi₂ $ λ i h, infi_le _ _) (le_infi $ λ ⟨i, h⟩, infi₂_le _ _) @[simp] lemma supr_exists {p : ι → Prop} {f : Exists p → α} : (⨆ x, f x) = ⨆ i h, f ⟨i, h⟩ := -@infi_exists (order_dual α) _ _ _ _ +@infi_exists αᵒᵈ _ _ _ _ lemma infi_and {p q : Prop} {s : p ∧ q → α} : infi s = ⨅ h₁ h₂, s ⟨h₁, h₂⟩ := le_antisymm (le_infi₂ $ λ i j, infi_le _ _) (le_infi $ λ ⟨i, h⟩, infi₂_le _ _) @@ -866,8 +856,7 @@ lemma infi_and' {p q : Prop} {s : p → q → α} : (⨅ (h₁ : p) (h₂ : q), s h₁ h₂) = ⨅ (h : p ∧ q), s h.1 h.2 := by { symmetry, exact infi_and } -lemma supr_and {p q : Prop} {s : p ∧ q → α} : supr s = ⨆ h₁ h₂, s ⟨h₁, h₂⟩ := -@infi_and (order_dual α) _ _ _ _ +lemma supr_and {p q : Prop} {s : p ∧ q → α} : supr s = ⨆ h₁ h₂, s ⟨h₁, h₂⟩ := @infi_and αᵒᵈ _ _ _ _ /-- The symmetric case of `supr_and`, useful for rewriting into a supremum over a conjunction -/ lemma supr_and' {p q : Prop} {s : p → q → α} : @@ -885,7 +874,7 @@ le_antisymm theorem supr_or {p q : Prop} {s : p ∨ q → α} : (⨆ x, s x) = (⨆ i, s (or.inl i)) ⊔ (⨆ j, s (or.inr j)) := -@infi_or (order_dual α) _ _ _ _ +@infi_or αᵒᵈ _ _ _ _ section @@ -906,7 +895,7 @@ supr_dite _ _ _ lemma infi_dite (f : Π i, p i → α) (g : Π i, ¬p i → α) : (⨅ i, if h : p i then f i h else g i h) = (⨅ i (h : p i), f i h) ⊓ (⨅ i (h : ¬ p i), g i h) := -supr_dite p (show Π i, p i → order_dual α, from f) g +supr_dite p (show Π i, p i → αᵒᵈ, from f) g lemma infi_ite (f g : ι → α) : (⨅ i, if p i then f i else g i) = (⨅ i (h : p i), f i) ⊓ (⨅ i (h : ¬ p i), g i) := @@ -918,13 +907,12 @@ lemma infi_range {g : β → α} {f : ι → β} : (⨅ b ∈ range f, g b) = by rw [← infi_subtype'', infi_range'] lemma supr_range {g : β → α} {f : ι → β} : (⨆ b ∈ range f, g b) = ⨆ i, g (f i) := -@infi_range (order_dual α) _ _ _ _ _ +@infi_range αᵒᵈ _ _ _ _ _ theorem Inf_image {s : set β} {f : β → α} : Inf (f '' s) = ⨅ a ∈ s, f a := by rw [← infi_subtype'', Inf_image'] -theorem Sup_image {s : set β} {f : β → α} : Sup (f '' s) = ⨆ a ∈ s, f a := -@Inf_image (order_dual α) _ _ _ _ +theorem Sup_image {s : set β} {f : β → α} : Sup (f '' s) = ⨆ a ∈ s, f a := @Inf_image αᵒᵈ _ _ _ _ /- ### supr and infi under set constructions @@ -959,19 +947,19 @@ by rw [(union_eq_self_of_subset_left h).symm, infi_union]; exact inf_le_left theorem supr_union {f : β → α} {s t : set β} : (⨆ x ∈ s ∪ t, f x) = (⨆ x ∈ s, f x) ⊔ (⨆ x ∈ t, f x) := -@infi_union (order_dual α) _ _ _ _ _ +@infi_union αᵒᵈ _ _ _ _ _ lemma supr_split (f : β → α) (p : β → Prop) : (⨆ i, f i) = (⨆ i (h : p i), f i) ⊔ (⨆ i (h : ¬ p i), f i) := -@infi_split (order_dual α) _ _ _ _ +@infi_split αᵒᵈ _ _ _ _ lemma supr_split_single (f : β → α) (i₀ : β) : (⨆ i, f i) = f i₀ ⊔ (⨆ i (h : i ≠ i₀), f i) := -@infi_split_single (order_dual α) _ _ _ _ +@infi_split_single αᵒᵈ _ _ _ _ theorem supr_le_supr_of_subset {f : β → α} {s t : set β} (h : s ⊆ t) : (⨆ x ∈ s, f x) ≤ (⨆ x ∈ t, f x) := -@infi_le_infi_of_subset (order_dual α) _ _ _ _ _ h +@infi_le_infi_of_subset αᵒᵈ _ _ _ _ _ h theorem infi_insert {f : β → α} {s : set β} {b : β} : (⨅ x ∈ insert b s, f x) = f b ⊓ (⨅ x ∈ s, f x) := @@ -988,7 +976,7 @@ theorem infi_pair {f : β → α} {a b : β} : (⨅ x ∈ ({a, b} : set β), f x by rw [infi_insert, infi_singleton] theorem supr_singleton {f : β → α} {b : β} : (⨆ x ∈ (singleton b : set β), f x) = f b := -@infi_singleton (order_dual α) _ _ _ _ +@infi_singleton αᵒᵈ _ _ _ _ theorem supr_pair {f : β → α} {a b : β} : (⨆ x ∈ ({a, b} : set β), f x) = f a ⊔ f b := by rw [supr_insert, supr_singleton] @@ -999,7 +987,7 @@ by rw [← Inf_image, ← Inf_image, ← image_comp] lemma supr_image {γ} {f : β → γ} {g : γ → α} {t : set β} : (⨆ c ∈ f '' t, g c) = (⨆ b ∈ t, g (f b)) := -@infi_image (order_dual α) _ _ _ _ _ _ +@infi_image αᵒᵈ _ _ _ _ _ _ theorem supr_extend_bot {e : ι → β} (he : injective e) (f : ι → α) : (⨆ j, extend e f ⊥ j) = ⨆ i, f i := @@ -1023,26 +1011,23 @@ theorem infi_of_empty' {α ι} [has_Inf α] [is_empty ι] (f : ι → α) : infi f = Inf (∅ : set α) := congr_arg Inf (range_eq_empty f) -theorem infi_of_empty [is_empty ι] (f : ι → α) : infi f = ⊤ := -@supr_of_empty (order_dual α) _ _ _ f +theorem infi_of_empty [is_empty ι] (f : ι → α) : infi f = ⊤ := @supr_of_empty αᵒᵈ _ _ _ f lemma supr_bool_eq {f : bool → α} : (⨆b:bool, f b) = f tt ⊔ f ff := by rw [supr, bool.range_eq, Sup_pair, sup_comm] -lemma infi_bool_eq {f : bool → α} : (⨅b:bool, f b) = f tt ⊓ f ff := -@supr_bool_eq (order_dual α) _ _ +lemma infi_bool_eq {f : bool → α} : (⨅b:bool, f b) = f tt ⊓ f ff := @supr_bool_eq αᵒᵈ _ _ lemma sup_eq_supr (x y : α) : x ⊔ y = ⨆ b : bool, cond b x y := by rw [supr_bool_eq, bool.cond_tt, bool.cond_ff] -lemma inf_eq_infi (x y : α) : x ⊓ y = ⨅ b : bool, cond b x y := -@sup_eq_supr (order_dual α) _ _ _ +lemma inf_eq_infi (x y : α) : x ⊓ y = ⨅ b : bool, cond b x y := @sup_eq_supr αᵒᵈ _ _ _ lemma is_glb_binfi {s : set β} {f : β → α} : is_glb (f '' s) (⨅ x ∈ s, f x) := by simpa only [range_comp, subtype.range_coe, infi_subtype'] using @is_glb_infi α s _ (f ∘ coe) theorem supr_subtype {p : ι → Prop} {f : subtype p → α} : (⨆ x, f x) = (⨆ i (h:p i), f ⟨i, h⟩) := -@infi_subtype (order_dual α) _ _ _ _ +@infi_subtype αᵒᵈ _ _ _ _ lemma supr_subtype' {p : ι → Prop} {f : ∀ i, p i → α} : (⨆ i (h : p i), f i h) = (⨆ x : subtype p, f x x.property) := @@ -1059,13 +1044,13 @@ theorem infi_sigma {p : β → Type*} {f : sigma p → α} : (⨅ x, f x) = (⨅ eq_of_forall_le_iff $ λ c, by simp only [le_infi_iff, sigma.forall] theorem supr_sigma {p : β → Type*} {f : sigma p → α} : (⨆ x, f x) = (⨆ i (h : p i), f ⟨i, h⟩) := -@infi_sigma (order_dual α) _ _ _ _ +@infi_sigma αᵒᵈ _ _ _ _ theorem infi_prod {γ : Type*} {f : β × γ → α} : (⨅ x, f x) = (⨅ i j, f (i, j)) := eq_of_forall_le_iff $ λ c, by simp only [le_infi_iff, prod.forall] theorem supr_prod {γ : Type*} {f : β × γ → α} : (⨆ x, f x) = (⨆ i j, f (i, j)) := -@infi_prod (order_dual α) _ _ _ _ +@infi_prod αᵒᵈ _ _ _ _ theorem infi_sum {γ : Type*} {f : β ⊕ γ → α} : (⨅ x, f x) = (⨅ i, f (sum.inl i)) ⊓ (⨅ j, f (sum.inr j)) := @@ -1073,7 +1058,7 @@ eq_of_forall_le_iff $ λ c, by simp only [le_inf_iff, le_infi_iff, sum.forall] theorem supr_sum {γ : Type*} {f : β ⊕ γ → α} : (⨆ x, f x) = (⨆ i, f (sum.inl i)) ⊔ (⨆ j, f (sum.inr j)) := -@infi_sum (order_dual α) _ _ _ _ +@infi_sum αᵒᵈ _ _ _ _ theorem supr_option (f : option β → α) : (⨆ o, f o) = f none ⊔ ⨆ b, f (option.some b) := @@ -1081,7 +1066,7 @@ eq_of_forall_ge_iff $ λ c, by simp only [supr_le_iff, sup_le_iff, option.forall theorem infi_option (f : option β → α) : (⨅ o, f o) = f none ⊓ ⨅ b, f (option.some b) := -@supr_option (order_dual α) _ _ _ +@supr_option αᵒᵈ _ _ _ /-- A version of `supr_option` useful for rewriting right-to-left. -/ lemma supr_option_elim (a : α) (f : β → α) : (⨆ o : option β, o.elim a f) = a ⊔ ⨆ b, f b := @@ -1089,7 +1074,7 @@ by simp [supr_option] /-- A version of `infi_option` useful for rewriting right-to-left. -/ lemma infi_option_elim (a : α) (f : β → α) : (⨅ o : option β, o.elim a f) = a ⊓ ⨅ b, f b := -@supr_option_elim (order_dual α) _ _ _ _ +@supr_option_elim αᵒᵈ _ _ _ _ /-- When taking the supremum of `f : ι → α`, the elements of `ι` on which `f` gives `⊥` can be dropped, without changing the result. -/ @@ -1109,7 +1094,7 @@ end /-- When taking the infimum of `f : ι → α`, the elements of `ι` on which `f` gives `⊤` can be dropped, without changing the result. -/ lemma infi_ne_top_subtype (f : ι → α) : (⨅ i : {i // f i ≠ ⊤}, f i) = ⨅ i, f i := -@supr_ne_bot_subtype (order_dual α) ι _ f +@supr_ne_bot_subtype αᵒᵈ ι _ f /-! ### `supr` and `infi` under `ℕ` @@ -1124,7 +1109,7 @@ begin end lemma infi_ge_eq_infi_nat_add {u : ℕ → α} (n : ℕ) : (⨅ i ≥ n, u i) = ⨅ i, u (i + n) := -@supr_ge_eq_supr_nat_add (order_dual α) _ _ _ +@supr_ge_eq_supr_nat_add αᵒᵈ _ _ _ lemma monotone.supr_nat_add {f : ℕ → α} (hf : monotone f) (k : ℕ) : (⨆ n, f (n + k)) = ⨆ n, f n := @@ -1148,7 +1133,7 @@ begin end lemma inf_infi_nat_succ (u : ℕ → α) : u 0 ⊓ (⨅ i, u (i + 1)) = ⨅ i, u i := -@sup_supr_nat_succ (order_dual α) _ u +@sup_supr_nat_succ αᵒᵈ _ u end @@ -1228,7 +1213,7 @@ by { change (∃ _, _) ↔ _, simp [-eq_iff_iff] } @[simp] lemma supr_apply {α : Type*} {β : α → Type*} {ι : Sort*} [Π i, has_Sup (β i)] {f : ι → Π a, β a} {a : α} : (⨆ i, f i) a = (⨆ i, f i a) := -@infi_apply α (λ i, order_dual (β i)) _ _ f a +@infi_apply α (λ i, (β i)ᵒᵈ) _ _ f a section complete_lattice variables [preorder α] [complete_lattice β] diff --git a/src/order/concept.lean b/src/order/concept.lean index c8d112fc01e15..0327b0d58dc01 100644 --- a/src/order/concept.lean +++ b/src/order/concept.lean @@ -255,7 +255,7 @@ instance : inhabited (concept α β r) := ⟨⊥⟩ @[simp] lemma swap_lt_swap_iff : c.swap < d.swap ↔ d < c := snd_ssubset_snd_iff /-- The dual of a concept lattice is isomorphic to the concept lattice of the dual context. -/ -@[simps] def swap_equiv : order_dual (concept α β r) ≃o concept β α (function.swap r) := +@[simps] def swap_equiv : (concept α β r)ᵒᵈ ≃o concept β α (function.swap r) := { to_fun := swap ∘ of_dual, inv_fun := to_dual ∘ swap, left_inv := swap_swap, diff --git a/src/order/conditionally_complete_lattice.lean b/src/order/conditionally_complete_lattice.lean index cbf266460a395..dc46cae9c8bf2 100644 --- a/src/order/conditionally_complete_lattice.lean +++ b/src/order/conditionally_complete_lattice.lean @@ -50,10 +50,10 @@ noncomputable instance {α : Type*} [has_Inf α] : has_Inf (with_top α) := ⟨λ S, if S ⊆ {⊤} then ⊤ else ↑(Inf (coe ⁻¹' S : set α))⟩ noncomputable instance {α : Type*} [has_Sup α] : has_Sup (with_bot α) := -⟨(@with_top.has_Inf (order_dual α) _).Inf⟩ +⟨(@with_top.has_Inf αᵒᵈ _).Inf⟩ noncomputable instance {α : Type*} [preorder α] [has_Inf α] : has_Inf (with_bot α) := -⟨(@with_top.has_Sup (order_dual α) _ _).Sup⟩ +⟨(@with_top.has_Sup αᵒᵈ _ _).Sup⟩ @[simp] theorem with_top.cInf_empty {α : Type*} [has_Inf α] : Inf (∅ : set (with_top α)) = ⊤ := @@ -182,8 +182,7 @@ end section order_dual -instance (α : Type*) [conditionally_complete_lattice α] : - conditionally_complete_lattice (order_dual α) := +instance (α : Type*) [conditionally_complete_lattice α] : conditionally_complete_lattice αᵒᵈ := { le_cSup := @conditionally_complete_lattice.cInf_le α _, cSup_le := @conditionally_complete_lattice.le_cInf α _, le_cInf := @conditionally_complete_lattice.cSup_le α _, @@ -193,7 +192,7 @@ instance (α : Type*) [conditionally_complete_lattice α] : ..order_dual.lattice α } instance (α : Type*) [conditionally_complete_linear_order α] : - conditionally_complete_linear_order (order_dual α) := + conditionally_complete_linear_order αᵒᵈ := { ..order_dual.conditionally_complete_lattice α, ..order_dual.linear_order α } @@ -254,7 +253,7 @@ is_glb_cInf (range_nonempty f) H lemma is_glb_cinfi_set {f : β → α} {s : set β} (H : bdd_below (f '' s)) (Hne : s.nonempty) : is_glb (f '' s) (⨅ i : s, f i) := -@is_lub_csupr_set (order_dual α) _ _ _ _ H Hne +@is_lub_csupr_set αᵒᵈ _ _ _ _ H Hne lemma is_lub.cSup_eq (H : is_lub s a) (ne : s.nonempty) : Sup s = a := (is_lub_cSup ne ⟨a, H.1⟩).unique H @@ -312,7 +311,7 @@ lemma not_mem_of_lt_cInf {x : α} {s : set α} (h : x < Inf s) (hs : bdd_below s λ hx, lt_irrefl _ (h.trans_le (cInf_le hs hx)) lemma not_mem_of_cSup_lt {x : α} {s : set α} (h : Sup s < x) (hs : bdd_above s) : x ∉ s := -@not_mem_of_lt_cInf (order_dual α) _ x s h hs +@not_mem_of_lt_cInf αᵒᵈ _ x s h hs /--Introduction rule to prove that `b` is the supremum of `s`: it suffices to check that `b` is larger than all elements of `s`, and that this is not the case of any `wb`. See `Inf_eq_of_forall_ge_of_forall_gt_exists_lt` for a version in complete lattices. -/ theorem cInf_eq_of_forall_ge_of_forall_gt_exists_lt (_ : s.nonempty) (_ : ∀a∈s, b ≤ a) (H : ∀w, b < w → (∃a∈s, a < w)) : Inf s = b := -@cSup_eq_of_forall_le_of_forall_lt_exists_gt (order_dual α) _ _ _ ‹_› ‹_› ‹_› +@cSup_eq_of_forall_le_of_forall_lt_exists_gt αᵒᵈ _ _ _ ‹_› ‹_› ‹_› /--b < Sup s when there is an element a in s with b < a, when s is bounded above. This is essentially an iff, except that the assumptions for the two implications are @@ -349,7 +348,7 @@ slightly different (one needs boundedness below for one direction, nonemptiness order for the other one), so we formulate separately the two implications, contrary to the complete_lattice case.-/ lemma cInf_lt_of_lt (_ : bdd_below s) (_ : a ∈ s) (_ : a < b) : Inf s < b := -@lt_cSup_of_lt (order_dual α) _ _ _ _ ‹_› ‹_› ‹_› +@lt_cSup_of_lt αᵒᵈ _ _ _ _ ‹_› ‹_› ‹_› /-- If all elements of a nonempty set `s` are less than or equal to all elements of a nonempty set `t`, then there exists an element between these sets. -/ @@ -387,7 +386,7 @@ theorem cSup_union (hs : bdd_above s) (sne : s.nonempty) (ht : bdd_above t) (tne that all sets are bounded below and nonempty.-/ theorem cInf_union (hs : bdd_below s) (sne : s.nonempty) (ht : bdd_below t) (tne : t.nonempty) : Inf (s ∪ t) = Inf s ⊓ Inf t := -@cSup_union (order_dual α) _ _ _ hs sne ht tne +@cSup_union αᵒᵈ _ _ _ hs sne ht tne /--The supremum of an intersection of two sets is bounded by the minimum of the suprema of each set, if all sets are bounded above and nonempty.-/ @@ -403,7 +402,7 @@ end infima of each set, if all sets are bounded below and nonempty.-/ theorem le_cInf_inter (_ : bdd_below s) (_ : bdd_below t) (hst : (s ∩ t).nonempty) : Inf s ⊔ Inf t ≤ Inf (s ∩ t) := -@cSup_inter_le (order_dual α) _ _ _ ‹_› ‹_› hst +@cSup_inter_le αᵒᵈ _ _ _ ‹_› ‹_› hst /-- The supremum of insert a s is the maximum of a and the supremum of s, if s is nonempty and bounded above.-/ @@ -413,7 +412,7 @@ theorem cSup_insert (hs : bdd_above s) (sne : s.nonempty) : Sup (insert a s) = a /-- The infimum of insert a s is the minimum of a and the infimum of s, if s is nonempty and bounded below.-/ theorem cInf_insert (hs : bdd_below s) (sne : s.nonempty) : Inf (insert a s) = a ⊓ Inf s := -@cSup_insert (order_dual α) _ _ _ hs sne +@cSup_insert αᵒᵈ _ _ _ hs sne @[simp] lemma cInf_Icc (h : a ≤ b) : Inf (Icc a b) = a := (is_glb_Icc h).cInf_eq (nonempty_Icc.2 h) @@ -478,41 +477,40 @@ lemma le_csupr_set {f : β → α} {s : set β} /--The indexed infimum of two functions are comparable if the functions are pointwise comparable-/ lemma cinfi_mono {f g : ι → α} (B : bdd_below (range f)) (H : ∀ x, f x ≤ g x) : infi f ≤ infi g := -@csupr_mono (order_dual α) _ _ _ _ B H +@csupr_mono αᵒᵈ _ _ _ _ B H /--The indexed minimum of a function is bounded below by a uniform lower bound-/ lemma le_cinfi [nonempty ι] {f : ι → α} {c : α} (H : ∀x, c ≤ f x) : c ≤ infi f := -@csupr_le (order_dual α) _ _ _ _ _ H +@csupr_le αᵒᵈ _ _ _ _ _ H /--The indexed infimum of a function is bounded above by the value taken at one point-/ lemma cinfi_le {f : ι → α} (H : bdd_below (range f)) (c : ι) : infi f ≤ f c := -@le_csupr (order_dual α) _ _ _ H c +@le_csupr αᵒᵈ _ _ _ H c lemma cinfi_le_of_le {f : ι → α} (H : bdd_below (range f)) (c : ι) (h : f c ≤ a) : infi f ≤ a := -@le_csupr_of_le (order_dual α) _ _ _ _ H c h +@le_csupr_of_le αᵒᵈ _ _ _ _ H c h lemma cinfi_set_le {f : β → α} {s : set β} (H : bdd_below (f '' s)) {c : β} (hc : c ∈ s) : (⨅ i : s, f i) ≤ f c := -@le_csupr_set (order_dual α) _ _ _ _ H _ hc +@le_csupr_set αᵒᵈ _ _ _ _ H _ hc @[simp] theorem csupr_const [hι : nonempty ι] {a : α} : (⨆ b:ι, a) = a := by rw [supr, range_const, cSup_singleton] -@[simp] theorem cinfi_const [hι : nonempty ι] {a : α} : (⨅ b:ι, a) = a := -@csupr_const (order_dual α) _ _ _ _ +@[simp] theorem cinfi_const [hι : nonempty ι] {a : α} : (⨅ b:ι, a) = a := @csupr_const αᵒᵈ _ _ _ _ @[simp] theorem supr_unique [unique ι] {s : ι → α} : (⨆ i, s i) = s default := have ∀ i, s i = s default := λ i, congr_arg s (unique.eq_default i), by simp only [this, csupr_const] @[simp] theorem infi_unique [unique ι] {s : ι → α} : (⨅ i, s i) = s default := -@supr_unique (order_dual α) _ _ _ _ +@supr_unique αᵒᵈ _ _ _ _ @[simp] lemma csupr_pos {p : Prop} {f : p → α} (hp : p) : (⨆ h : p, f h) = f hp := by haveI := unique_prop hp; exact supr_unique @[simp] lemma cinfi_pos {p : Prop} {f : p → α} (hp : p) : (⨅ h : p, f h) = f hp := -@csupr_pos (order_dual α) _ _ _ hp +@csupr_pos αᵒᵈ _ _ _ hp lemma csupr_set {s : set β} {f : β → α} : (⨆ x : s, f x) = Sup (f '' s) := begin @@ -523,8 +521,7 @@ begin simp_rw [subtype.coe_mk, exists_prop], end -lemma cinfi_set {s : set β} {f : β → α} : (⨅ x : s, f x) = Inf (f '' s) := -@csupr_set (order_dual α) _ _ _ _ +lemma cinfi_set {s : set β} {f : β → α} : (⨅ x : s, f x) = Inf (f '' s) := @csupr_set αᵒᵈ _ _ _ _ /--Introduction rule to prove that `b` is the supremum of `f`: it suffices to check that `b` is larger than `f i` for all `i`, and that this is not the case of any `wb`. See `infi_eq_of_forall_ge_of_forall_gt_exists_lt` for a version in complete lattices. -/ theorem cinfi_eq_of_forall_ge_of_forall_gt_exists_lt [nonempty ι] {f : ι → α} (h₁ : ∀ i, b ≤ f i) (h₂ : ∀ w, b < w → (∃ i, f i < w)) : (⨅ (i : ι), f i) = b := -@csupr_eq_of_forall_le_of_forall_lt_exists_gt (order_dual α) _ _ _ _ ‹_› ‹_› ‹_› +@csupr_eq_of_forall_le_of_forall_lt_exists_gt αᵒᵈ _ _ _ _ ‹_› ‹_› ‹_› /-- Nested intervals lemma: if `f` is a monotone sequence, `g` is an antitone sequence, and `f n ≤ g n` for all `n`, then `⨆ n, f n` belongs to all the intervals `[f n, g n]`. -/ @@ -601,25 +598,25 @@ lemma finset.nonempty.cSup_eq_max' {s : finset α} (h : s.nonempty) : Sup ↑s = eq_of_forall_ge_iff $ λ a, (cSup_le_iff s.bdd_above h.to_set).trans (s.max'_le_iff h).symm lemma finset.nonempty.cInf_eq_min' {s : finset α} (h : s.nonempty) : Inf ↑s = s.min' h := -@finset.nonempty.cSup_eq_max' (order_dual α) _ s h +@finset.nonempty.cSup_eq_max' αᵒᵈ _ s h lemma finset.nonempty.cSup_mem {s : finset α} (h : s.nonempty) : Sup (s : set α) ∈ s := by { rw h.cSup_eq_max', exact s.max'_mem _ } lemma finset.nonempty.cInf_mem {s : finset α} (h : s.nonempty) : Inf (s : set α) ∈ s := -@finset.nonempty.cSup_mem (order_dual α) _ _ h +@finset.nonempty.cSup_mem αᵒᵈ _ _ h lemma set.nonempty.cSup_mem (h : s.nonempty) (hs : finite s) : Sup s ∈ s := by { lift s to finset α using hs, exact finset.nonempty.cSup_mem h } lemma set.nonempty.cInf_mem (h : s.nonempty) (hs : finite s) : Inf s ∈ s := -@set.nonempty.cSup_mem (order_dual α) _ _ h hs +@set.nonempty.cSup_mem αᵒᵈ _ _ h hs lemma set.finite.cSup_lt_iff (hs : finite s) (h : s.nonempty) : Sup s < a ↔ ∀ x ∈ s, x < a := ⟨λ h x hx, (le_cSup hs.bdd_above hx).trans_lt h, λ H, H _ $ h.cSup_mem hs⟩ lemma set.finite.lt_cInf_iff (hs : finite s) (h : s.nonempty) : a < Inf s ↔ ∀ x ∈ s, a < x := -@set.finite.cSup_lt_iff (order_dual α) _ _ _ hs h +@set.finite.cSup_lt_iff αᵒᵈ _ _ _ hs h /-- When b < Sup s, there is an element a in s with b < a, if s is nonempty and the order is a linear order. -/ @@ -640,7 +637,7 @@ let ⟨_, ⟨i, rfl⟩, h⟩ := exists_lt_of_lt_cSup (range_nonempty f) h in ⟨ /--When Inf s < b, there is an element a in s with a < b, if s is nonempty and the order is a linear order.-/ lemma exists_lt_of_cInf_lt (hs : s.nonempty) (hb : Inf s < b) : ∃a∈s, a < b := -@exists_lt_of_lt_cSup (order_dual α) _ _ _ hs hb +@exists_lt_of_lt_cSup αᵒᵈ _ _ _ hs hb /-- Indexed version of the above lemma `exists_lt_of_cInf_lt` @@ -648,7 +645,7 @@ When `infi f < a`, there is an element `i` such that `f i < a`. -/ lemma exists_lt_of_cinfi_lt [nonempty ι] {f : ι → α} (h : infi f < a) : (∃i, f i < a) := -@exists_lt_of_lt_csupr (order_dual α) _ _ _ _ _ h +@exists_lt_of_lt_csupr αᵒᵈ _ _ _ _ _ h open function variables [is_well_order α (<)] @@ -859,11 +856,11 @@ cSup_le (nonempty.image f hs) (h_mono.mem_upper_bounds_image hB) lemma cInf_image_le {s : set α} {c : α} (hcs : c ∈ s) (h_bdd : bdd_below s) : Inf (f '' s) ≤ f c := -@le_cSup_image (order_dual α) (order_dual β) _ _ _ (λ x y hxy, h_mono hxy) _ _ hcs h_bdd +@le_cSup_image αᵒᵈ βᵒᵈ _ _ _ (λ x y hxy, h_mono hxy) _ _ hcs h_bdd lemma le_cInf_image {s : set α} (hs : s.nonempty) {B : α} (hB: B ∈ lower_bounds s) : f B ≤ Inf (f '' s) := -@cSup_image_le (order_dual α) (order_dual β) _ _ _ (λ x y hxy, h_mono hxy) _ hs _ hB +@cSup_image_le αᵒᵈ βᵒᵈ _ _ _ (λ x y hxy, h_mono hxy) _ hs _ hB end monotone @@ -977,7 +974,7 @@ end lemma inf'_eq_cInf_image [conditionally_complete_lattice β] (s : finset α) (H) (f : α → β) : s.inf' H f = Inf (f '' s) := -@sup'_eq_cSup_image _ (order_dual β) _ _ _ _ +@sup'_eq_cSup_image _ βᵒᵈ _ _ _ _ lemma sup'_id_eq_cSup [conditionally_complete_lattice α] (s : finset α) (H) : s.sup' H id = Sup s := @@ -985,7 +982,7 @@ by rw [sup'_eq_cSup_image s H, set.image_id] lemma inf'_id_eq_cInf [conditionally_complete_lattice α] (s : finset α) (H) : s.inf' H id = Inf s := -@sup'_id_eq_cSup (order_dual α) _ _ _ +@sup'_id_eq_cSup αᵒᵈ _ _ _ end finset @@ -1025,10 +1022,10 @@ gives a conditionally complete lattice -/ noncomputable instance with_bot.conditionally_complete_lattice {α : Type*} [conditionally_complete_lattice α] : conditionally_complete_lattice (with_bot α) := -{ le_cSup := (@with_top.conditionally_complete_lattice (order_dual α) _).cInf_le, - cSup_le := (@with_top.conditionally_complete_lattice (order_dual α) _).le_cInf, - cInf_le := (@with_top.conditionally_complete_lattice (order_dual α) _).le_cSup, - le_cInf := (@with_top.conditionally_complete_lattice (order_dual α) _).cSup_le, +{ le_cSup := (@with_top.conditionally_complete_lattice αᵒᵈ _).cInf_le, + cSup_le := (@with_top.conditionally_complete_lattice αᵒᵈ _).le_cInf, + cInf_le := (@with_top.conditionally_complete_lattice αᵒᵈ _).le_cSup, + le_cInf := (@with_top.conditionally_complete_lattice αᵒᵈ _).cSup_le, ..with_bot.lattice, ..with_bot.has_Sup, ..with_bot.has_Inf } @@ -1081,7 +1078,7 @@ inv_mul_le_iff_le_mul.mp $ le_cinfi $ λ hi, inv_mul_le_iff_le_mul.mpr $ H _ @[to_additive] lemma mul_csupr_le [covariant_class α α (*) (≤)] {a : α} {g : α} {h : ι → α} (H : ∀ j, g * h j ≤ a) : g * supr h ≤ a := -@le_mul_cinfi (order_dual α) _ _ _ _ _ _ _ _ H +@le_mul_cinfi αᵒᵈ _ _ _ _ _ _ _ _ H @[to_additive] lemma le_cinfi_mul [covariant_class α α (function.swap (*)) (≤)] {a : α} {g : ι → α} {h : α} @@ -1091,7 +1088,7 @@ mul_inv_le_iff_le_mul.mp $ le_cinfi $ λ gi, mul_inv_le_iff_le_mul.mpr $ H _ @[to_additive] lemma csupr_mul_le [covariant_class α α (function.swap (*)) (≤)] {a : α} {g : ι → α} {h : α} (H : ∀ i, g i * h ≤ a) : supr g * h ≤ a := -@le_cinfi_mul (order_dual α) _ _ _ _ _ _ _ _ H +@le_cinfi_mul αᵒᵈ _ _ _ _ _ _ _ _ H @[to_additive] lemma le_cinfi_mul_cinfi [covariant_class α α (*) (≤)] [covariant_class α α (function.swap (*)) (≤)] diff --git a/src/order/cover.lean b/src/order/cover.lean index 0fd3a7fa85f0b..bcb6c46c22f2a 100644 --- a/src/order/cover.lean +++ b/src/order/cover.lean @@ -84,7 +84,7 @@ lemma set.ord_connected.apply_wcovby_apply_iff (f : α ↪o β) (h : (range f).o @[simp] lemma to_dual_wcovby_to_dual_iff : to_dual b ⩿ to_dual a ↔ a ⩿ b := and_congr_right' $ forall_congr $ λ c, forall_swap -@[simp] lemma of_dual_wcovby_of_dual_iff {a b : order_dual α} : +@[simp] lemma of_dual_wcovby_of_dual_iff {a b : αᵒᵈ} : of_dual a ⩿ of_dual b ↔ b ⩿ a := and_congr_right' $ forall_congr $ λ c, forall_swap @@ -148,8 +148,7 @@ lemma densely_ordered_iff_forall_not_covby : densely_ordered α ↔ ∀ a b : α @[simp] lemma to_dual_covby_to_dual_iff : to_dual b ⋖ to_dual a ↔ a ⋖ b := and_congr_right' $ forall_congr $ λ c, forall_swap -@[simp] lemma of_dual_covby_of_dual_iff {a b : order_dual α} : - of_dual a ⋖ of_dual b ↔ b ⋖ a := +@[simp] lemma of_dual_covby_of_dual_iff {a b : αᵒᵈ} : of_dual a ⋖ of_dual b ↔ b ⋖ a := and_congr_right' $ forall_congr $ λ c, forall_swap alias to_dual_covby_to_dual_iff ↔ _ covby.to_dual diff --git a/src/order/directed.lean b/src/order/directed.lean index c7ace46f1073b..11f44fa57d220 100644 --- a/src/order/directed.lean +++ b/src/order/directed.lean @@ -122,12 +122,10 @@ directed_of (≤) a b lemma exists_le_le [has_le α] [is_directed α (swap (≤))] (a b : α) : ∃ c, c ≤ a ∧ c ≤ b := directed_of (swap (≤)) a b -instance order_dual.is_directed_ge [has_le α] [is_directed α (≤)] : - is_directed (order_dual α) (swap (≤)) := +instance order_dual.is_directed_ge [has_le α] [is_directed α (≤)] : is_directed αᵒᵈ (swap (≤)) := by assumption -instance order_dual.is_directed_le [has_le α] [is_directed α (swap (≤))] : - is_directed (order_dual α) (≤) := +instance order_dual.is_directed_le [has_le α] [is_directed α (swap (≤))] : is_directed αᵒᵈ (≤) := by assumption section preorder diff --git a/src/order/filter/at_top_bot.lean b/src/order/filter/at_top_bot.lean index 1c6947f521a19..f28d73a903a54 100644 --- a/src/order/filter/at_top_bot.lean +++ b/src/order/filter/at_top_bot.lean @@ -56,7 +56,7 @@ lemma disjoint_at_bot_principal_Ioi [preorder α] (x : α) : disjoint at_bot ( disjoint_of_disjoint_of_mem (Iic_disjoint_Ioi le_rfl) (Iic_mem_at_bot x) (mem_principal_self _) lemma disjoint_at_top_principal_Iio [preorder α] (x : α) : disjoint at_top (𝓟 (Iio x)) := -@disjoint_at_bot_principal_Ioi (order_dual α) _ _ +@disjoint_at_bot_principal_Ioi αᵒᵈ _ _ lemma disjoint_at_top_principal_Iic [preorder α] [no_max_order α] (x : α) : disjoint at_top (𝓟 (Iic x)) := @@ -64,7 +64,7 @@ disjoint_of_disjoint_of_mem (Iic_disjoint_Ioi le_rfl).symm (Ioi_mem_at_top x) (m lemma disjoint_at_bot_principal_Ici [preorder α] [no_min_order α] (x : α) : disjoint at_bot (𝓟 (Ici x)) := -@disjoint_at_top_principal_Iic (order_dual α) _ _ _ +@disjoint_at_top_principal_Iic αᵒᵈ _ _ _ lemma disjoint_at_bot_at_top [partial_order α] [nontrivial α] : disjoint (at_bot : filter α) at_top := @@ -91,13 +91,11 @@ lemma at_top_basis' [semilattice_sup α] (a : α) : ⟨λ ⟨x, _, hx⟩, ⟨x ⊔ a, le_sup_right, λ y hy, hx (le_trans le_sup_left hy)⟩, λ ⟨x, _, hx⟩, ⟨x, trivial, hx⟩⟩⟩ -lemma at_bot_basis [nonempty α] [semilattice_inf α] : - (@at_bot α _).has_basis (λ _, true) Iic := -@at_top_basis (order_dual α) _ _ +lemma at_bot_basis [nonempty α] [semilattice_inf α] : (@at_bot α _).has_basis (λ _, true) Iic := +@at_top_basis αᵒᵈ _ _ -lemma at_bot_basis' [semilattice_inf α] (a : α) : - (@at_bot α _).has_basis (λ x, x ≤ a) Iic := -@at_top_basis' (order_dual α) _ _ +lemma at_bot_basis' [semilattice_inf α] (a : α) : (@at_bot α _).has_basis (λ x, x ≤ a) Iic := +@at_top_basis' αᵒᵈ _ _ @[instance] lemma at_top_ne_bot [nonempty α] [semilattice_sup α] : ne_bot (at_top : filter α) := @@ -105,7 +103,7 @@ at_top_basis.ne_bot_iff.2 $ λ a _, nonempty_Ici @[instance] lemma at_bot_ne_bot [nonempty α] [semilattice_inf α] : ne_bot (at_bot : filter α) := -@at_top_ne_bot (order_dual α) _ _ +@at_top_ne_bot αᵒᵈ _ _ @[simp] lemma mem_at_top_sets [nonempty α] [semilattice_sup α] {s : set α} : @@ -115,7 +113,7 @@ at_top_basis.mem_iff.trans $ exists_congr $ λ _, exists_const _ @[simp] lemma mem_at_bot_sets [nonempty α] [semilattice_inf α] {s : set α} : s ∈ (at_bot : filter α) ↔ ∃a:α, ∀b≤a, b ∈ s := -@mem_at_top_sets (order_dual α) _ _ _ +@mem_at_top_sets αᵒᵈ _ _ _ @[simp] lemma eventually_at_top [semilattice_sup α] [nonempty α] {p : α → Prop} : @@ -173,7 +171,7 @@ le_antisymm (le_pure_iff.2 $ (eventually_ge_at_top ⊤).mono $ λ b, top_unique) (le_infi $ λ b, le_principal_iff.2 le_top) lemma order_bot.at_bot_eq (α) [partial_order α] [order_bot α] : (at_bot : filter α) = pure ⊥ := -@order_top.at_top_eq (order_dual α) _ _ +@order_top.at_top_eq αᵒᵈ _ _ @[nontriviality] lemma subsingleton.at_top_eq (α) [subsingleton α] [preorder α] : (at_top : filter α) = ⊤ := @@ -186,7 +184,7 @@ end @[nontriviality] lemma subsingleton.at_bot_eq (α) [subsingleton α] [preorder α] : (at_bot : filter α) = ⊤ := -@subsingleton.at_top_eq (order_dual α) _ _ +@subsingleton.at_top_eq αᵒᵈ _ _ lemma tendsto_at_top_pure [partial_order α] [order_top α] (f : α → β) : tendsto f at_top (pure $ f ⊤) := @@ -194,7 +192,7 @@ lemma tendsto_at_top_pure [partial_order α] [order_top α] (f : α → β) : lemma tendsto_at_bot_pure [partial_order α] [order_bot α] (f : α → β) : tendsto f at_bot (pure $ f ⊥) := -@tendsto_at_top_pure (order_dual α) _ _ _ _ +@tendsto_at_top_pure αᵒᵈ _ _ _ _ lemma eventually.exists_forall_of_at_top [semilattice_sup α] [nonempty α] {p : α → Prop} (h : ∀ᶠ x in at_top, p x) : ∃ a, ∀ b ≥ a, p b := @@ -210,7 +208,7 @@ by simp [at_top_basis.frequently_iff] lemma frequently_at_bot [semilattice_inf α] [nonempty α] {p : α → Prop} : (∃ᶠ x in at_bot, p x) ↔ (∀ a, ∃ b ≤ a, p b) := -@frequently_at_top (order_dual α) _ _ _ +@frequently_at_top αᵒᵈ _ _ _ lemma frequently_at_top' [semilattice_sup α] [nonempty α] [no_max_order α] {p : α → Prop} : (∃ᶠ x in at_top, p x) ↔ (∀ a, ∃ b > a, p b) := @@ -218,7 +216,7 @@ by simp [at_top_basis_Ioi.frequently_iff] lemma frequently_at_bot' [semilattice_inf α] [nonempty α] [no_min_order α] {p : α → Prop} : (∃ᶠ x in at_bot, p x) ↔ (∀ a, ∃ b < a, p b) := -@frequently_at_top' (order_dual α) _ _ _ _ +@frequently_at_top' αᵒᵈ _ _ _ _ lemma frequently.forall_exists_of_at_top [semilattice_sup α] [nonempty α] {p : α → Prop} (h : ∃ᶠ x in at_top, p x) : ∀ a, ∃ b ≥ a, p b := @@ -234,7 +232,7 @@ lemma map_at_top_eq [nonempty α] [semilattice_sup α] {f : α → β} : lemma map_at_bot_eq [nonempty α] [semilattice_inf α] {f : α → β} : at_bot.map f = (⨅a, 𝓟 $ f '' {a' | a' ≤ a}) := -@map_at_top_eq (order_dual α) _ _ _ _ +@map_at_top_eq αᵒᵈ _ _ _ _ lemma tendsto_at_top [preorder β] {m : α → β} {f : filter α} : tendsto m f at_top ↔ (∀b, ∀ᶠ a in f, b ≤ m a) := @@ -242,7 +240,7 @@ by simp only [at_top, tendsto_infi, tendsto_principal, mem_Ici] lemma tendsto_at_bot [preorder β] {m : α → β} {f : filter α} : tendsto m f at_bot ↔ (∀b, ∀ᶠ a in f, m a ≤ b) := -@tendsto_at_top α (order_dual β) _ m f +@tendsto_at_top α βᵒᵈ _ m f lemma tendsto_at_top_mono' [preorder β] (l : filter α) ⦃f₁ f₂ : α → β⦄ (h : f₁ ≤ᶠ[l] f₂) : tendsto f₁ l at_top → tendsto f₂ l at_top := @@ -251,7 +249,7 @@ assume h₁, tendsto_at_top.2 $ λ b, mp_mem (tendsto_at_top.1 h₁ b) lemma tendsto_at_bot_mono' [preorder β] (l : filter α) ⦃f₁ f₂ : α → β⦄ (h : f₁ ≤ᶠ[l] f₂) : tendsto f₂ l at_bot → tendsto f₁ l at_bot := -@tendsto_at_top_mono' _ (order_dual β) _ _ _ _ h +@tendsto_at_top_mono' _ βᵒᵈ _ _ _ _ h lemma tendsto_at_top_mono [preorder β] {l : filter α} {f g : α → β} (h : ∀ n, f n ≤ g n) : tendsto f l at_top → tendsto g l at_top := @@ -259,7 +257,7 @@ tendsto_at_top_mono' l $ eventually_of_forall h lemma tendsto_at_bot_mono [preorder β] {l : filter α} {f g : α → β} (h : ∀ n, f n ≤ g n) : tendsto g l at_bot → tendsto f l at_bot := -@tendsto_at_top_mono _ (order_dual β) _ _ _ _ h +@tendsto_at_top_mono _ βᵒᵈ _ _ _ _ h /-! ### Sequences @@ -271,7 +269,7 @@ by simp_rw [inf_ne_bot_iff_frequently_left, frequently_map, frequently_at_top]; lemma inf_map_at_bot_ne_bot_iff [semilattice_inf α] [nonempty α] {F : filter β} {u : α → β} : ne_bot (F ⊓ (map u at_bot)) ↔ ∀ U ∈ F, ∀ N, ∃ n ≤ N, u n ∈ U := -@inf_map_at_top_ne_bot_iff (order_dual α) _ _ _ _ _ +@inf_map_at_top_ne_bot_iff αᵒᵈ _ _ _ _ _ lemma extraction_of_frequently_at_top' {P : ℕ → Prop} (h : ∀ N, ∃ n > N, P n) : ∃ φ : ℕ → ℕ, strict_mono φ ∧ ∀ n, P (φ n) := @@ -327,7 +325,7 @@ end @[nolint ge_or_gt] -- see Note [nolint_ge] lemma exists_le_of_tendsto_at_bot [semilattice_sup α] [preorder β] {u : α → β} (h : tendsto u at_top at_bot) : ∀ a b, ∃ a' ≥ a, u a' ≤ b := -@exists_le_of_tendsto_at_top _ (order_dual β) _ _ _ h +@exists_le_of_tendsto_at_top _ βᵒᵈ _ _ _ h lemma exists_lt_of_tendsto_at_top [semilattice_sup α] [preorder β] [no_max_order β] {u : α → β} (h : tendsto u at_top at_top) (a : α) (b : β) : ∃ a' ≥ a, b < u a' := @@ -340,7 +338,7 @@ end @[nolint ge_or_gt] -- see Note [nolint_ge] lemma exists_lt_of_tendsto_at_bot [semilattice_sup α] [preorder β] [no_min_order β] {u : α → β} (h : tendsto u at_top at_bot) : ∀ a b, ∃ a' ≥ a, u a' < b := -@exists_lt_of_tendsto_at_top _ (order_dual β) _ _ _ _ h +@exists_lt_of_tendsto_at_top _ βᵒᵈ _ _ _ _ h /-- If `u` is a sequence which is unbounded above, @@ -376,7 +374,7 @@ then after any point, it reaches a value strictly smaller than all previous valu @[nolint ge_or_gt] -- see Note [nolint_ge] lemma low_scores [linear_order β] [no_min_order β] {u : ℕ → β} (hu : tendsto u at_top at_bot) : ∀ N, ∃ n ≥ N, ∀ k < n, u n < u k := -@high_scores (order_dual β) _ _ _ hu +@high_scores βᵒᵈ _ _ _ hu /-- If `u` is a sequence which is unbounded above, @@ -392,7 +390,7 @@ then it `frequently` reaches a value strictly smaller than all previous values. -/ lemma frequently_low_scores [linear_order β] [no_min_order β] {u : ℕ → β} (hu : tendsto u at_top at_bot) : ∃ᶠ n in at_top, ∀ k < n, u n < u k := -@frequently_high_scores (order_dual β) _ _ _ hu +@frequently_high_scores βᵒᵈ _ _ _ hu lemma strict_mono_subseq_of_tendsto_at_top {β : Type*} [linear_order β] [no_max_order β] @@ -419,7 +417,7 @@ tendsto_at_top_mono' l (hf.mono (λ x, le_add_of_nonneg_left)) hg lemma tendsto_at_bot_add_nonpos_left' (hf : ∀ᶠ x in l, f x ≤ 0) (hg : tendsto g l at_bot) : tendsto (λ x, f x + g x) l at_bot := -@tendsto_at_top_add_nonneg_left' _ (order_dual β) _ _ _ _ hf hg +@tendsto_at_top_add_nonneg_left' _ βᵒᵈ _ _ _ _ hf hg lemma tendsto_at_top_add_nonneg_left (hf : ∀ x, 0 ≤ f x) (hg : tendsto g l at_top) : tendsto (λ x, f x + g x) l at_top := @@ -427,7 +425,7 @@ tendsto_at_top_add_nonneg_left' (eventually_of_forall hf) hg lemma tendsto_at_bot_add_nonpos_left (hf : ∀ x, f x ≤ 0) (hg : tendsto g l at_bot) : tendsto (λ x, f x + g x) l at_bot := -@tendsto_at_top_add_nonneg_left _ (order_dual β) _ _ _ _ hf hg +@tendsto_at_top_add_nonneg_left _ βᵒᵈ _ _ _ _ hf hg lemma tendsto_at_top_add_nonneg_right' (hf : tendsto f l at_top) (hg : ∀ᶠ x in l, 0 ≤ g x) : tendsto (λ x, f x + g x) l at_top := @@ -435,7 +433,7 @@ tendsto_at_top_mono' l (monotone_mem (λ x, le_add_of_nonneg_right) hg) hf lemma tendsto_at_bot_add_nonpos_right' (hf : tendsto f l at_bot) (hg : ∀ᶠ x in l, g x ≤ 0) : tendsto (λ x, f x + g x) l at_bot := -@tendsto_at_top_add_nonneg_right' _ (order_dual β) _ _ _ _ hf hg +@tendsto_at_top_add_nonneg_right' _ βᵒᵈ _ _ _ _ hf hg lemma tendsto_at_top_add_nonneg_right (hf : tendsto f l at_top) (hg : ∀ x, 0 ≤ g x) : tendsto (λ x, f x + g x) l at_top := @@ -443,7 +441,7 @@ tendsto_at_top_add_nonneg_right' hf (eventually_of_forall hg) lemma tendsto_at_bot_add_nonpos_right (hf : tendsto f l at_bot) (hg : ∀ x, g x ≤ 0) : tendsto (λ x, f x + g x) l at_bot := -@tendsto_at_top_add_nonneg_right _ (order_dual β) _ _ _ _ hf hg +@tendsto_at_top_add_nonneg_right _ βᵒᵈ _ _ _ _ hf hg lemma tendsto_at_top_add (hf : tendsto f l at_top) (hg : tendsto g l at_top) : tendsto (λ x, f x + g x) l at_top := @@ -451,7 +449,7 @@ tendsto_at_top_add_nonneg_left' (tendsto_at_top.mp hf 0) hg lemma tendsto_at_bot_add (hf : tendsto f l at_bot) (hg : tendsto g l at_bot) : tendsto (λ x, f x + g x) l at_bot := -@tendsto_at_top_add _ (order_dual β) _ _ _ _ hf hg +@tendsto_at_top_add _ βᵒᵈ _ _ _ _ hf hg lemma tendsto.nsmul_at_top (hf : tendsto f l at_top) {n : ℕ} (hn : 0 < n) : tendsto (λ x, n • f x) l at_top := @@ -462,7 +460,7 @@ calc y ≤ f x : hy lemma tendsto.nsmul_at_bot (hf : tendsto f l at_bot) {n : ℕ} (hn : 0 < n) : tendsto (λ x, n • f x) l at_bot := -@tendsto.nsmul_at_top α (order_dual β) _ l f hf n hn +@tendsto.nsmul_at_top α βᵒᵈ _ l f hf n hn lemma tendsto_bit0_at_top : tendsto bit0 (at_top : filter β) at_top := tendsto_at_top_add tendsto_id tendsto_id @@ -482,7 +480,7 @@ tendsto_at_top.2 $ assume b, (tendsto_at_top.1 hf (C + b)).mono (λ x, le_of_add lemma tendsto_at_bot_of_add_const_left (C : β) (hf : tendsto (λ x, C + f x) l at_bot) : tendsto f l at_bot := -@tendsto_at_top_of_add_const_left _ (order_dual β) _ _ _ C hf +@tendsto_at_top_of_add_const_left _ βᵒᵈ _ _ _ C hf lemma tendsto_at_top_of_add_const_right (C : β) (hf : tendsto (λ x, f x + C) l at_top) : tendsto f l at_top := @@ -490,7 +488,7 @@ tendsto_at_top.2 $ assume b, (tendsto_at_top.1 hf (b + C)).mono (λ x, le_of_add lemma tendsto_at_bot_of_add_const_right (C : β) (hf : tendsto (λ x, f x + C) l at_bot) : tendsto f l at_bot := -@tendsto_at_top_of_add_const_right _ (order_dual β) _ _ _ C hf +@tendsto_at_top_of_add_const_right _ βᵒᵈ _ _ _ C hf lemma tendsto_at_top_of_add_bdd_above_left' (C) (hC : ∀ᶠ x in l, f x ≤ C) (h : tendsto (λ x, f x + g x) l at_top) : @@ -501,7 +499,7 @@ tendsto_at_top_of_add_const_left C lemma tendsto_at_bot_of_add_bdd_below_left' (C) (hC : ∀ᶠ x in l, C ≤ f x) (h : tendsto (λ x, f x + g x) l at_bot) : tendsto g l at_bot := -@tendsto_at_top_of_add_bdd_above_left' _ (order_dual β) _ _ _ _ C hC h +@tendsto_at_top_of_add_bdd_above_left' _ βᵒᵈ _ _ _ _ C hC h lemma tendsto_at_top_of_add_bdd_above_left (C) (hC : ∀ x, f x ≤ C) : tendsto (λ x, f x + g x) l at_top → tendsto g l at_top := @@ -509,7 +507,7 @@ tendsto_at_top_of_add_bdd_above_left' C (univ_mem' hC) lemma tendsto_at_bot_of_add_bdd_below_left (C) (hC : ∀ x, C ≤ f x) : tendsto (λ x, f x + g x) l at_bot → tendsto g l at_bot := -@tendsto_at_top_of_add_bdd_above_left _ (order_dual β) _ _ _ _ C hC +@tendsto_at_top_of_add_bdd_above_left _ βᵒᵈ _ _ _ _ C hC lemma tendsto_at_top_of_add_bdd_above_right' (C) (hC : ∀ᶠ x in l, g x ≤ C) (h : tendsto (λ x, f x + g x) l at_top) : @@ -520,7 +518,7 @@ tendsto_at_top_of_add_const_right C lemma tendsto_at_bot_of_add_bdd_below_right' (C) (hC : ∀ᶠ x in l, C ≤ g x) (h : tendsto (λ x, f x + g x) l at_bot) : tendsto f l at_bot := -@tendsto_at_top_of_add_bdd_above_right' _ (order_dual β) _ _ _ _ C hC h +@tendsto_at_top_of_add_bdd_above_right' _ βᵒᵈ _ _ _ _ C hC h lemma tendsto_at_top_of_add_bdd_above_right (C) (hC : ∀ x, g x ≤ C) : tendsto (λ x, f x + g x) l at_top → tendsto f l at_top := @@ -528,7 +526,7 @@ tendsto_at_top_of_add_bdd_above_right' C (univ_mem' hC) lemma tendsto_at_bot_of_add_bdd_below_right (C) (hC : ∀ x, C ≤ g x) : tendsto (λ x, f x + g x) l at_bot → tendsto f l at_bot := -@tendsto_at_top_of_add_bdd_above_right _ (order_dual β) _ _ _ _ C hC +@tendsto_at_top_of_add_bdd_above_right _ βᵒᵈ _ _ _ _ C hC end ordered_cancel_add_comm_monoid @@ -543,7 +541,7 @@ lemma tendsto_at_top_add_left_of_le' (C : β) (hf : ∀ᶠ x in l, C ≤ f x) (h lemma tendsto_at_bot_add_left_of_ge' (C : β) (hf : ∀ᶠ x in l, f x ≤ C) (hg : tendsto g l at_bot) : tendsto (λ x, f x + g x) l at_bot := -@tendsto_at_top_add_left_of_le' _ (order_dual β) _ _ _ _ C hf hg +@tendsto_at_top_add_left_of_le' _ βᵒᵈ _ _ _ _ C hf hg lemma tendsto_at_top_add_left_of_le (C : β) (hf : ∀ x, C ≤ f x) (hg : tendsto g l at_top) : tendsto (λ x, f x + g x) l at_top := @@ -551,7 +549,7 @@ tendsto_at_top_add_left_of_le' l C (univ_mem' hf) hg lemma tendsto_at_bot_add_left_of_ge (C : β) (hf : ∀ x, f x ≤ C) (hg : tendsto g l at_bot) : tendsto (λ x, f x + g x) l at_bot := -@tendsto_at_top_add_left_of_le _ (order_dual β) _ _ _ _ C hf hg +@tendsto_at_top_add_left_of_le _ βᵒᵈ _ _ _ _ C hf hg lemma tendsto_at_top_add_right_of_le' (C : β) (hf : tendsto f l at_top) (hg : ∀ᶠ x in l, C ≤ g x) : tendsto (λ x, f x + g x) l at_top := @@ -560,7 +558,7 @@ lemma tendsto_at_top_add_right_of_le' (C : β) (hf : tendsto f l at_top) (hg : lemma tendsto_at_bot_add_right_of_ge' (C : β) (hf : tendsto f l at_bot) (hg : ∀ᶠ x in l, g x ≤ C) : tendsto (λ x, f x + g x) l at_bot := -@tendsto_at_top_add_right_of_le' _ (order_dual β) _ _ _ _ C hf hg +@tendsto_at_top_add_right_of_le' _ βᵒᵈ _ _ _ _ C hf hg lemma tendsto_at_top_add_right_of_le (C : β) (hf : tendsto f l at_top) (hg : ∀ x, C ≤ g x) : tendsto (λ x, f x + g x) l at_top := @@ -568,7 +566,7 @@ tendsto_at_top_add_right_of_le' l C hf (univ_mem' hg) lemma tendsto_at_bot_add_right_of_ge (C : β) (hf : tendsto f l at_bot) (hg : ∀ x, g x ≤ C) : tendsto (λ x, f x + g x) l at_bot := -@tendsto_at_top_add_right_of_le _ (order_dual β) _ _ _ _ C hf hg +@tendsto_at_top_add_right_of_le _ βᵒᵈ _ _ _ _ C hf hg lemma tendsto_at_top_add_const_left (C : β) (hf : tendsto f l at_top) : tendsto (λ x, C + f x) l at_top := @@ -576,7 +574,7 @@ tendsto_at_top_add_left_of_le' l C (univ_mem' $ λ _, le_refl C) hf lemma tendsto_at_bot_add_const_left (C : β) (hf : tendsto f l at_bot) : tendsto (λ x, C + f x) l at_bot := -@tendsto_at_top_add_const_left _ (order_dual β) _ _ _ C hf +@tendsto_at_top_add_const_left _ βᵒᵈ _ _ _ C hf lemma tendsto_at_top_add_const_right (C : β) (hf : tendsto f l at_top) : tendsto (λ x, f x + C) l at_top := @@ -584,7 +582,7 @@ tendsto_at_top_add_right_of_le' l C hf (univ_mem' $ λ _, le_refl C) lemma tendsto_at_bot_add_const_right (C : β) (hf : tendsto f l at_bot) : tendsto (λ x, f x + C) l at_bot := -@tendsto_at_top_add_const_right _ (order_dual β) _ _ _ C hf +@tendsto_at_top_add_const_right _ βᵒᵈ _ _ _ C hf lemma tendsto_neg_at_top_at_bot : tendsto (has_neg.neg : β → β) at_top at_bot := begin @@ -593,14 +591,14 @@ begin end lemma tendsto_neg_at_bot_at_top : tendsto (has_neg.neg : β → β) at_bot at_top := -@tendsto_neg_at_top_at_bot (order_dual β) _ +@tendsto_neg_at_top_at_bot βᵒᵈ _ lemma tendsto_at_top_iff_tends_to_neg_at_bot : tendsto f l at_top ↔ tendsto (-f) l at_bot := have hf : f = has_neg.neg ∘ -f, { ext, simp, }, ⟨tendsto_neg_at_top_at_bot.comp, λ h, hf.symm ▸ tendsto_neg_at_bot_at_top.comp h⟩ lemma tendsto_at_bot_iff_tends_to_neg_at_top : tendsto f l at_bot ↔ tendsto (-f) l at_top := -@tendsto_at_top_iff_tends_to_neg_at_bot α (order_dual β) _ l f +@tendsto_at_top_iff_tends_to_neg_at_bot α βᵒᵈ _ l f end ordered_group @@ -827,7 +825,7 @@ by simp only [tendsto_def, mem_at_top_sets]; refl lemma tendsto_at_bot' [nonempty α] [semilattice_inf α] {f : α → β} {l : filter β} : tendsto f at_bot l ↔ (∀s ∈ l, ∃a, ∀b≤a, f b ∈ s) := -@tendsto_at_top' (order_dual α) _ _ _ _ _ +@tendsto_at_top' αᵒᵈ _ _ _ _ _ theorem tendsto_at_top_principal [nonempty β] [semilattice_sup β] {f : β → α} {s : set α} : tendsto f at_top (𝓟 s) ↔ ∃N, ∀n≥N, f n ∈ s := @@ -835,7 +833,7 @@ by rw [tendsto_iff_comap, comap_principal, le_principal_iff, mem_at_top_sets]; r theorem tendsto_at_bot_principal [nonempty β] [semilattice_inf β] {f : β → α} {s : set α} : tendsto f at_bot (𝓟 s) ↔ ∃N, ∀n≤N, f n ∈ s := -@tendsto_at_top_principal _ (order_dual β) _ _ _ _ +@tendsto_at_top_principal _ βᵒᵈ _ _ _ _ /-- A function `f` grows to `+∞` independent of an order-preserving embedding `e`. -/ lemma tendsto_at_top_at_top [nonempty α] [semilattice_sup α] [preorder β] {f : α → β} : @@ -844,15 +842,15 @@ iff.trans tendsto_infi $ forall_congr $ assume b, tendsto_at_top_principal lemma tendsto_at_top_at_bot [nonempty α] [semilattice_sup α] [preorder β] {f : α → β} : tendsto f at_top at_bot ↔ ∀ (b : β), ∃ (i : α), ∀ (a : α), i ≤ a → f a ≤ b := -@tendsto_at_top_at_top α (order_dual β) _ _ _ f +@tendsto_at_top_at_top α βᵒᵈ _ _ _ f lemma tendsto_at_bot_at_top [nonempty α] [semilattice_inf α] [preorder β] {f : α → β} : tendsto f at_bot at_top ↔ ∀ (b : β), ∃ (i : α), ∀ (a : α), a ≤ i → b ≤ f a := -@tendsto_at_top_at_top (order_dual α) β _ _ _ f +@tendsto_at_top_at_top αᵒᵈ β _ _ _ f lemma tendsto_at_bot_at_bot [nonempty α] [semilattice_inf α] [preorder β] {f : α → β} : tendsto f at_bot at_bot ↔ ∀ (b : β), ∃ (i : α), ∀ (a : α), a ≤ i → f a ≤ b := -@tendsto_at_top_at_top (order_dual α) (order_dual β) _ _ _ f +@tendsto_at_top_at_top αᵒᵈ βᵒᵈ _ _ _ f lemma tendsto_at_top_at_top_of_monotone [preorder α] [preorder β] {f : α → β} (hf : monotone f) (h : ∀ b, ∃ a, b ≤ f a) : @@ -893,7 +891,7 @@ le_antisymm lemma comap_embedding_at_bot [preorder β] [preorder γ] {e : β → γ} (hm : ∀ b₁ b₂, e b₁ ≤ e b₂ ↔ b₁ ≤ b₂) (hu : ∀c, ∃b, e b ≤ c) : comap e at_bot = at_bot := -@comap_embedding_at_top (order_dual β) (order_dual γ) _ _ e (function.swap hm) hu +@comap_embedding_at_top βᵒᵈ γᵒᵈ _ _ e (function.swap hm) hu lemma tendsto_at_top_embedding [preorder β] [preorder γ] {f : α → β} {e : β → γ} {l : filter α} @@ -906,7 +904,7 @@ lemma tendsto_at_bot_embedding [preorder β] [preorder γ] {f : α → β} {e : β → γ} {l : filter α} (hm : ∀b₁ b₂, e b₁ ≤ e b₂ ↔ b₁ ≤ b₂) (hu : ∀c, ∃b, e b ≤ c) : tendsto (e ∘ f) l at_bot ↔ tendsto f l at_bot := -@tendsto_at_top_embedding α (order_dual β) (order_dual γ) _ _ f e l (function.swap hm) hu +@tendsto_at_top_embedding α βᵒᵈ γᵒᵈ _ _ f e l (function.swap hm) hu lemma tendsto_finset_range : tendsto finset.range at_top at_top := finset.range_mono.tendsto_at_top_at_top finset.exists_nat_subset_range @@ -959,7 +957,7 @@ end lemma prod_at_bot_at_bot_eq {β₁ β₂ : Type*} [semilattice_inf β₁] [semilattice_inf β₂] : (at_bot : filter β₁) ×ᶠ (at_bot : filter β₂) = (at_bot : filter (β₁ × β₂)) := -@prod_at_top_at_top_eq (order_dual β₁) (order_dual β₂) _ _ +@prod_at_top_at_top_eq β₁ᵒᵈ β₂ᵒᵈ _ _ lemma prod_map_at_top_eq {α₁ α₂ β₁ β₂ : Type*} [semilattice_sup β₁] [semilattice_sup β₂] (u₁ : β₁ → α₁) (u₂ : β₂ → α₂) : @@ -969,7 +967,7 @@ by rw [prod_map_map_eq, prod_at_top_at_top_eq, prod.map_def] lemma prod_map_at_bot_eq {α₁ α₂ β₁ β₂ : Type*} [semilattice_inf β₁] [semilattice_inf β₂] (u₁ : β₁ → α₁) (u₂ : β₂ → α₂) : (map u₁ at_bot) ×ᶠ (map u₂ at_bot) = map (prod.map u₁ u₂) at_bot := -@prod_map_at_top_eq _ _ (order_dual β₁) (order_dual β₂) _ _ _ _ +@prod_map_at_top_eq _ _ β₁ᵒᵈ β₂ᵒᵈ _ _ _ _ lemma tendsto.subseq_mem {F : filter α} {V : ℕ → set α} (h : ∀ n, V n ∈ F) {u : ℕ → α} (hu : tendsto u at_top F) : ∃ φ : ℕ → ℕ, strict_mono φ ∧ ∀ n, u (φ n) ∈ V n := @@ -1036,7 +1034,7 @@ end lemma eventually_at_bot_curry [semilattice_inf α] [semilattice_inf β] {p : α × β → Prop} (hp : ∀ᶠ (x : α × β) in filter.at_bot, p x) : ∀ᶠ k in at_bot, ∀ᶠ l in at_bot, p (k, l) := -@eventually_at_top_curry (order_dual α) (order_dual β) _ _ _ hp +@eventually_at_top_curry αᵒᵈ βᵒᵈ _ _ _ hp /-- A function `f` maps upwards closed sets (at_top sets) to upwards closed sets when it is a Galois insertion. The Galois "insertion" and "connection" is weakened to only require it to be an @@ -1056,7 +1054,7 @@ end lemma map_at_bot_eq_of_gc [semilattice_inf α] [semilattice_inf β] {f : α → β} (g : β → α) (b' : β) (hf : monotone f) (gc : ∀a, ∀b≤b', b ≤ f a ↔ g b ≤ a) (hgi : ∀b≤b', f (g b) ≤ b) : map f at_bot = at_bot := -@map_at_top_eq_of_gc (order_dual α) (order_dual β) _ _ _ _ _ hf.dual gc hgi +@map_at_top_eq_of_gc αᵒᵈ βᵒᵈ _ _ _ _ _ hf.dual gc hgi lemma map_coe_at_top_of_Ici_subset [semilattice_sup α] {a : α} {s : set α} (h : Ici a ⊆ s) : map (coe : s → α) at_top = at_top := @@ -1109,25 +1107,24 @@ by rw [← map_coe_Ici_at_top a, comap_map subtype.coe_injective] order. -/ @[simp] lemma map_coe_Iio_at_bot [semilattice_inf α] [no_min_order α] (a : α) : map (coe : Iio a → α) at_bot = at_bot := -@map_coe_Ioi_at_top (order_dual α) _ _ _ +@map_coe_Ioi_at_top αᵒᵈ _ _ _ /-- The `at_bot` filter for an open interval `Iio a` comes from the `at_bot` filter in the ambient order. -/ -lemma at_bot_Iio_eq [semilattice_inf α] (a : α) : - at_bot = comap (coe : Iio a → α) at_bot := -@at_top_Ioi_eq (order_dual α) _ _ +lemma at_bot_Iio_eq [semilattice_inf α] (a : α) : at_bot = comap (coe : Iio a → α) at_bot := +@at_top_Ioi_eq αᵒᵈ _ _ /-- The `at_bot` filter for an open interval `Iic a` comes from the `at_bot` filter in the ambient order. -/ @[simp] lemma map_coe_Iic_at_bot [semilattice_inf α] (a : α) : map (coe : Iic a → α) at_bot = at_bot := -@map_coe_Ici_at_top (order_dual α) _ _ +@map_coe_Ici_at_top αᵒᵈ _ _ /-- The `at_bot` filter for an open interval `Iic a` comes from the `at_bot` filter in the ambient order. -/ lemma at_bot_Iic_eq [semilattice_inf α] (a : α) : at_bot = comap (coe : Iic a → α) at_bot := -@at_top_Ici_eq (order_dual α) _ _ +@at_top_Ici_eq αᵒᵈ _ _ lemma tendsto_Ioi_at_top [semilattice_sup α] {a : α} {f : β → Ioi a} {l : filter β} : @@ -1224,7 +1221,7 @@ below, then `tendsto u at_bot at_bot`. -/ lemma tendsto_at_bot_at_bot_of_monotone' [preorder ι] [linear_order α] {u : ι → α} (h : monotone u) (H : ¬bdd_below (range u)) : tendsto u at_bot at_bot := -@tendsto_at_top_at_top_of_monotone' (order_dual ι) (order_dual α) _ _ _ h.dual H +@tendsto_at_top_at_top_of_monotone' ιᵒᵈ αᵒᵈ _ _ _ h.dual H lemma unbounded_of_tendsto_at_top [nonempty α] [semilattice_sup α] [preorder β] [no_max_order β] {f : α → β} (h : tendsto f at_top at_top) : @@ -1241,17 +1238,17 @@ end lemma unbounded_of_tendsto_at_bot [nonempty α] [semilattice_sup α] [preorder β] [no_min_order β] {f : α → β} (h : tendsto f at_top at_bot) : ¬ bdd_below (range f) := -@unbounded_of_tendsto_at_top _ (order_dual β) _ _ _ _ _ h +@unbounded_of_tendsto_at_top _ βᵒᵈ _ _ _ _ _ h lemma unbounded_of_tendsto_at_top' [nonempty α] [semilattice_inf α] [preorder β] [no_max_order β] {f : α → β} (h : tendsto f at_bot at_top) : ¬ bdd_above (range f) := -@unbounded_of_tendsto_at_top (order_dual α) _ _ _ _ _ _ h +@unbounded_of_tendsto_at_top αᵒᵈ _ _ _ _ _ _ h lemma unbounded_of_tendsto_at_bot' [nonempty α] [semilattice_inf α] [preorder β] [no_min_order β] {f : α → β} (h : tendsto f at_bot at_bot) : ¬ bdd_below (range f) := -@unbounded_of_tendsto_at_top (order_dual α) (order_dual β) _ _ _ _ _ h +@unbounded_of_tendsto_at_top αᵒᵈ βᵒᵈ _ _ _ _ _ h /-- If a monotone function `u : ι → α` tends to `at_top` along *some* non-trivial filter `l`, then it tends to `at_top` along `at_top`. -/ @@ -1265,7 +1262,7 @@ it tends to `at_bot` along `at_bot`. -/ lemma tendsto_at_bot_of_monotone_of_filter [preorder ι] [preorder α] {l : filter ι} {u : ι → α} (h : monotone u) [ne_bot l] (hu : tendsto u l at_bot) : tendsto u at_bot at_bot := -@tendsto_at_top_of_monotone_of_filter (order_dual ι) (order_dual α) _ _ _ _ h.dual _ hu +@tendsto_at_top_of_monotone_of_filter ιᵒᵈ αᵒᵈ _ _ _ _ h.dual _ hu lemma tendsto_at_top_of_monotone_of_subseq [preorder ι] [preorder α] {u : ι → α} {φ : ι' → ι} (h : monotone u) {l : filter ι'} [ne_bot l] diff --git a/src/order/filter/basic.lean b/src/order/filter/basic.lean index 68c9530e6aa07..748cf89a8400b 100644 --- a/src/order/filter/basic.lean +++ b/src/order/filter/basic.lean @@ -351,7 +351,7 @@ show u ∈ (filter.mk_of_closure s hs).sets ↔ u ∈ (generate s).sets, from hs /-- Galois insertion from sets of sets into filters. -/ def gi_generate (α : Type*) : - @galois_insertion (set (set α)) (order_dual (filter α)) _ _ filter.generate filter.sets := + @galois_insertion (set (set α)) (filter α)ᵒᵈ _ _ filter.generate filter.sets := { gc := λ s f, sets_iff_generate, le_l_u := λ f u h, generate_sets.basic h, choice := λ s hs, filter.mk_of_closure s (le_antisymm hs $ sets_iff_generate.1 $ le_rfl), diff --git a/src/order/filter/cofinite.lean b/src/order/filter/cofinite.lean index 33f72e0f88abd..38f8908f2dec6 100644 --- a/src/order/filter/cofinite.lean +++ b/src/order/filter/cofinite.lean @@ -143,12 +143,12 @@ let ⟨a₀, _, ha₀⟩ := hf.exists_within_forall_le univ_nonempty in ⟨a₀, lemma filter.tendsto.exists_within_forall_ge [linear_order β] {s : set α} (hs : s.nonempty) {f : α → β} (hf : filter.tendsto f filter.cofinite filter.at_bot) : ∃ a₀ ∈ s, ∀ a ∈ s, f a ≤ f a₀ := -@filter.tendsto.exists_within_forall_le _ (order_dual β) _ _ hs _ hf +@filter.tendsto.exists_within_forall_le _ βᵒᵈ _ _ hs _ hf lemma filter.tendsto.exists_forall_ge [nonempty α] [linear_order β] {f : α → β} (hf : tendsto f cofinite at_bot) : ∃ a₀, ∀ a, f a ≤ f a₀ := -@filter.tendsto.exists_forall_le _ (order_dual β) _ _ _ hf +@filter.tendsto.exists_forall_le _ βᵒᵈ _ _ _ hf /-- For an injective function `f`, inverse images of finite sets are finite. See also `filter.comap_cofinite_le` and `function.injective.comap_cofinite_eq`. -/ diff --git a/src/order/filter/extr.lean b/src/order/filter/extr.lean index 160c0bec9ac22..937e748bcbce8 100644 --- a/src/order/filter/extr.lean +++ b/src/order/filter/extr.lean @@ -529,7 +529,7 @@ lemma filter.eventually_eq.is_max_filter_iff {α β : Type*} [preorder β] {f g lemma filter.eventually_le.is_min_filter {α β : Type*} [preorder β] {f g : α → β} {a : α} {l : filter α} (hle : f ≤ᶠ[l] g) (hfga : f a = g a) (h : is_min_filter f l a) : is_min_filter g l a := -@filter.eventually_le.is_max_filter _ (order_dual β) _ _ _ _ _ hle hfga h +@filter.eventually_le.is_max_filter _ βᵒᵈ _ _ _ _ _ hle hfga h lemma is_min_filter.congr {α β : Type*} [preorder β] {f g : α → β} {a : α} {l : filter α} (h : is_min_filter f l a) (heq : f =ᶠ[l] g) (hfga : f a = g a) : @@ -568,8 +568,7 @@ begin exact csupr_eq_of_forall_le_of_forall_lt_exists_gt (λ x, h x.prop) (λ w hw, ⟨⟨x₀, hx₀⟩, hw⟩), end -lemma is_min_on.infi_eq (hx₀ : x₀ ∈ s) (h : is_min_on f s x₀) : - (⨅ x : s, f x) = f x₀ := -@is_max_on.supr_eq (order_dual α) β _ _ _ _ hx₀ h +lemma is_min_on.infi_eq (hx₀ : x₀ ∈ s) (h : is_min_on f s x₀) : (⨅ x : s, f x) = f x₀ := +@is_max_on.supr_eq αᵒᵈ β _ _ _ _ hx₀ h end conditionally_complete_linear_order diff --git a/src/order/fixed_points.lean b/src/order/fixed_points.lean index cb7fecc3fda1c..5ce5fa03dd506 100644 --- a/src/order/fixed_points.lean +++ b/src/order/fixed_points.lean @@ -137,10 +137,8 @@ begin ... = a : ha end -lemma gfp_gfp (h : α →o α →o α) : - gfp (gfp.comp h) = gfp h.on_diag := -@lfp_lfp (order_dual α) _ $ (order_hom.dual_iso (order_dual α) - (order_dual α)).symm.to_order_embedding.to_order_hom.comp h.dual +lemma gfp_gfp (h : α →o α →o α) : gfp (gfp.comp h) = gfp h.on_diag := +@lfp_lfp αᵒᵈ _ $ (order_hom.dual_iso αᵒᵈ αᵒᵈ).symm.to_order_embedding.to_order_hom.comp h.dual end eqn diff --git a/src/order/grade.lean b/src/order/grade.lean index 1aa5a116184b6..9a8f467c7cb7f 100644 --- a/src/order/grade.lean +++ b/src/order/grade.lean @@ -14,7 +14,7 @@ This file defines graded orders, also known as ranked orders. A `𝕆`-graded order is an order `α` equipped with a distinguished "grade" function `α → 𝕆` which should be understood as giving the "height" of the elements. Usual graded orders are `ℕ`-graded, -cograded orders are `order_dual ℕ`-graded, but we can also grade by `ℤ`, and polytopes are naturally +cograded orders are `ℕᵒᵈ`-graded, but we can also grade by `ℤ`, and polytopes are naturally `fin n`-graded. Visually, `grade ℕ a` is the height of `a` in the Hasse diagram of `α`. @@ -178,26 +178,26 @@ instance preorder.to_grade_bounded_order : grade_bounded_order α α := /-! #### Dual -/ -instance [grade_order 𝕆 α] : grade_order (order_dual 𝕆) (order_dual α) := +instance [grade_order 𝕆 α] : grade_order 𝕆ᵒᵈ αᵒᵈ := { grade := to_dual ∘ grade 𝕆 ∘ of_dual, grade_strict_mono := grade_strict_mono.dual, covby_grade := λ a b h, (h.of_dual.grade _).to_dual } -instance [grade_max_order 𝕆 α] : grade_min_order (order_dual 𝕆) (order_dual α) := +instance [grade_max_order 𝕆 α] : grade_min_order 𝕆ᵒᵈ αᵒᵈ := { is_min_grade := λ _, is_max.grade _, ..order_dual.grade_order } -instance [grade_min_order 𝕆 α] : grade_max_order (order_dual 𝕆) (order_dual α) := +instance [grade_min_order 𝕆 α] : grade_max_order 𝕆ᵒᵈ αᵒᵈ := { is_max_grade := λ _, is_min.grade _, ..order_dual.grade_order } -instance [grade_bounded_order 𝕆 α] : grade_bounded_order (order_dual 𝕆) (order_dual α) := +instance [grade_bounded_order 𝕆 α] : grade_bounded_order 𝕆ᵒᵈ αᵒᵈ := { ..order_dual.grade_min_order, ..order_dual.grade_max_order } @[simp] lemma grade_to_dual [grade_order 𝕆 α] (a : α) : - grade (order_dual 𝕆) (to_dual a) = to_dual (grade 𝕆 a) := rfl -@[simp] lemma grade_of_dual [grade_order 𝕆 α] (a : order_dual α) : - grade 𝕆 (of_dual a) = of_dual (grade (order_dual 𝕆) a) := rfl + grade 𝕆ᵒᵈ (to_dual a) = to_dual (grade 𝕆 a) := rfl +@[simp] lemma grade_of_dual [grade_order 𝕆 α] (a : αᵒᵈ) : + grade 𝕆 (of_dual a) = of_dual (grade 𝕆ᵒᵈ a) := rfl /-! #### Lifting a graded order -/ diff --git a/src/order/hom/basic.lean b/src/order/hom/basic.lean index 1c646e6b3e071..5db3beb839b61 100644 --- a/src/order/hom/basic.lean +++ b/src/order/hom/basic.lean @@ -51,11 +51,9 @@ because the more bundled version usually does not work with dot notation. `α →o Π i, π i`; * `order_hom.pi_iso`: order isomorphism between `α →o Π i, π i` and `Π i, α →o π i`; * `order_hom.subtyle.val`: embedding `subtype.val : subtype p → α` as a bundled monotone map; -* `order_hom.dual`: reinterpret a monotone map `α →o β` as a monotone map - `order_dual α →o order_dual β`; -* `order_hom.dual_iso`: order isomorphism between `α →o β` and - `order_dual (order_dual α →o order_dual β)`; -* `order_iso.compl`: order isomorphism `α ≃o order_dual α` given by taking complements in a +* `order_hom.dual`: reinterpret a monotone map `α →o β` as a monotone map `αᵒᵈ →o βᵒᵈ`; +* `order_hom.dual_iso`: order isomorphism between `α →o β` and `(αᵒᵈ →o βᵒᵈ)ᵒᵈ`; +* `order_iso.compl`: order isomorphism `α ≃o αᵒᵈ` given by taking complements in a boolean algebra; We also define two functions to convert other bundled maps to `α →o β`: @@ -347,7 +345,7 @@ lemma order_hom_eq_id [subsingleton α] (g : α →o α) : g = order_hom.id := subsingleton.elim _ _ /-- Reinterpret a bundled monotone function as a monotone function between dual orders. -/ -@[simps] protected def dual : (α →o β) ≃ (order_dual α →o order_dual β) := +@[simps] protected def dual : (α →o β) ≃ (αᵒᵈ →o βᵒᵈ) := { to_fun := λ f, ⟨order_dual.to_dual ∘ f ∘ order_dual.of_dual, f.mono.dual⟩, inv_fun := λ f, ⟨order_dual.of_dual ∘ f ∘ order_dual.to_dual, f.mono.dual⟩, left_inv := λ f, ext _ _ rfl, @@ -357,13 +355,11 @@ subsingleton.elim _ _ @[simp] lemma dual_comp (g : β →o γ) (f : α →o β) : (g.comp f).dual = g.dual.comp f.dual := rfl @[simp] lemma symm_dual_id : order_hom.dual.symm order_hom.id = (order_hom.id : α →o α) := rfl -@[simp] lemma symm_dual_comp (g : order_dual β →o order_dual γ) - (f : order_dual α →o order_dual β) : +@[simp] lemma symm_dual_comp (g : βᵒᵈ →o γᵒᵈ) (f : αᵒᵈ →o βᵒᵈ) : order_hom.dual.symm (g.comp f) = (order_hom.dual.symm g).comp (order_hom.dual.symm f) := rfl /-- `order_hom.dual` as an order isomorphism. -/ -def dual_iso (α β : Type*) [preorder α] [preorder β] : - (α →o β) ≃o order_dual (order_dual α →o order_dual β) := +def dual_iso (α β : Type*) [preorder α] [preorder β] : (α →o β) ≃o (αᵒᵈ →o βᵒᵈ)ᵒᵈ := { to_equiv := order_hom.dual.trans order_dual.to_dual, map_rel_iff' := λ f g, iff.rfl } @@ -412,7 +408,7 @@ protected theorem is_well_order [is_well_order β (<)] : is_well_order α (<) := f.lt_embedding.is_well_order /-- An order embedding is also an order embedding between dual orders. -/ -protected def dual : order_dual α ↪o order_dual β := +protected def dual : αᵒᵈ ↪o βᵒᵈ := ⟨f.to_embedding, λ a b, f.map_rel_iff⟩ /-- @@ -579,7 +575,7 @@ def prod_comm : (α × β) ≃o (β × α) := variables (α) /-- The order isomorphism between a type and its double dual. -/ -def dual_dual : α ≃o order_dual (order_dual α) := refl α +def dual_dual : α ≃o αᵒᵈᵒᵈ := refl α @[simp] lemma coe_dual_dual : ⇑(dual_dual α) = to_dual ∘ to_dual := rfl @[simp] lemma coe_dual_dual_symm : ⇑(dual_dual α).symm = of_dual ∘ of_dual := rfl @@ -587,9 +583,7 @@ def dual_dual : α ≃o order_dual (order_dual α) := refl α variables {α} @[simp] lemma dual_dual_apply (a : α) : dual_dual α a = to_dual (to_dual a) := rfl - -@[simp] lemma dual_dual_symm_apply (a : order_dual (order_dual α)) : - (dual_dual α).symm a = of_dual (of_dual a) := rfl +@[simp] lemma dual_dual_symm_apply (a : αᵒᵈᵒᵈ) : (dual_dual α).symm a = of_dual (of_dual a) := rfl end has_le @@ -715,8 +709,8 @@ lemma order_iso_of_surjective_self_symm_apply (b : β) : end strict_mono /-- An order isomorphism is also an order isomorphism between dual orders. -/ -protected def order_iso.dual [has_le α] [has_le β] (f : α ≃o β) : - order_dual α ≃o order_dual β := ⟨f.to_equiv, λ _ _, f.map_rel_iff⟩ +protected def order_iso.dual [has_le α] [has_le β] (f : α ≃o β) : αᵒᵈ ≃o βᵒᵈ := +⟨f.to_equiv, λ _ _, f.map_rel_iff⟩ section lattice_isos @@ -774,8 +768,7 @@ f.dual.map_inf x y namespace with_bot /-- Taking the dual then adding `⊥` is the same as adding `⊤` then taking the dual. -/ -protected def to_dual_top [has_le α] : with_bot (order_dual α) ≃o order_dual (with_top α) := -order_iso.refl _ +protected def to_dual_top [has_le α] : with_bot αᵒᵈ ≃o (with_top α)ᵒᵈ := order_iso.refl _ @[simp] lemma to_dual_top_coe [has_le α] (a : α) : with_bot.to_dual_top ↑(to_dual a) = to_dual (a : with_top α) := rfl @@ -787,8 +780,7 @@ end with_bot namespace with_top /-- Taking the dual then adding `⊤` is the same as adding `⊥` then taking the dual. -/ -protected def to_dual_bot [has_le α] : with_top (order_dual α) ≃o order_dual (with_bot α) := -order_iso.refl _ +protected def to_dual_bot [has_le α] : with_top αᵒᵈ ≃o (with_bot α)ᵒᵈ := order_iso.refl _ @[simp] lemma to_dual_bot_coe [has_le α] (a : α) : with_top.to_dual_bot ↑(to_dual a) = to_dual (a : with_bot α) := rfl @@ -873,7 +865,7 @@ variables (α) [boolean_algebra α] /-- Taking complements as an order isomorphism to the order dual. -/ @[simps] -def order_iso.compl : α ≃o order_dual α := +def order_iso.compl : α ≃o αᵒᵈ := { to_fun := order_dual.to_dual ∘ compl, inv_fun := compl ∘ order_dual.of_dual, left_inv := compl_compl, diff --git a/src/order/hom/bounded.lean b/src/order/hom/bounded.lean index d3c8d3bf916f9..e35a40a93b5e4 100644 --- a/src/order/hom/bounded.lean +++ b/src/order/hom/bounded.lean @@ -436,7 +436,7 @@ namespace top_hom variables [has_le α] [order_top α] [has_le β] [order_top β] [has_le γ] [order_top γ] /-- Reinterpret a top homomorphism as a bot homomorphism between the dual lattices. -/ -@[simps] protected def dual : top_hom α β ≃ bot_hom (order_dual α) (order_dual β) := +@[simps] protected def dual : top_hom α β ≃ bot_hom αᵒᵈ βᵒᵈ := { to_fun := λ f, ⟨f, f.map_top'⟩, inv_fun := λ f, ⟨f, f.map_bot'⟩, left_inv := λ f, top_hom.ext $ λ _, rfl, @@ -447,8 +447,7 @@ variables [has_le α] [order_top α] [has_le β] [order_top β] [has_le γ] [ord (g.comp f).dual = g.dual.comp f.dual := rfl @[simp] lemma symm_dual_id : top_hom.dual.symm (bot_hom.id _) = top_hom.id α := rfl -@[simp] lemma symm_dual_comp (g : bot_hom (order_dual β) (order_dual γ)) - (f : bot_hom (order_dual α) (order_dual β)) : +@[simp] lemma symm_dual_comp (g : bot_hom βᵒᵈ γᵒᵈ) (f : bot_hom αᵒᵈ βᵒᵈ) : top_hom.dual.symm (g.comp f) = (top_hom.dual.symm g).comp (top_hom.dual.symm f) := rfl end top_hom @@ -457,7 +456,7 @@ namespace bot_hom variables [has_le α] [order_bot α] [has_le β] [order_bot β] [has_le γ] [order_bot γ] /-- Reinterpret a bot homomorphism as a top homomorphism between the dual lattices. -/ -@[simps] protected def dual : bot_hom α β ≃ top_hom (order_dual α) (order_dual β) := +@[simps] protected def dual : bot_hom α β ≃ top_hom αᵒᵈ βᵒᵈ := { to_fun := λ f, ⟨f, f.map_bot'⟩, inv_fun := λ f, ⟨f, f.map_top'⟩, left_inv := λ f, bot_hom.ext $ λ _, rfl, @@ -468,8 +467,7 @@ variables [has_le α] [order_bot α] [has_le β] [order_bot β] [has_le γ] [ord (g.comp f).dual = g.dual.comp f.dual := rfl @[simp] lemma symm_dual_id : bot_hom.dual.symm (top_hom.id _) = bot_hom.id α := rfl -@[simp] lemma symm_dual_comp (g : top_hom (order_dual β) (order_dual γ)) - (f : top_hom (order_dual α) (order_dual β)) : +@[simp] lemma symm_dual_comp (g : top_hom βᵒᵈ γᵒᵈ) (f : top_hom αᵒᵈ βᵒᵈ) : bot_hom.dual.symm (g.comp f) = (bot_hom.dual.symm g).comp (bot_hom.dual.symm f) := rfl end bot_hom @@ -480,8 +478,7 @@ variables [preorder α] [bounded_order α] [preorder β] [bounded_order β] [pre /-- Reinterpret a bounded order homomorphism as a bounded order homomorphism between the dual orders. -/ -@[simps] protected def dual : - bounded_order_hom α β ≃ bounded_order_hom (order_dual α) (order_dual β) := +@[simps] protected def dual : bounded_order_hom α β ≃ bounded_order_hom αᵒᵈ βᵒᵈ := { to_fun := λ f, ⟨f.to_order_hom.dual, f.map_bot', f.map_top'⟩, inv_fun := λ f, ⟨order_hom.dual.symm f.to_order_hom, f.map_bot', f.map_top'⟩, left_inv := λ f, ext $ λ a, rfl, @@ -493,8 +490,7 @@ orders. -/ @[simp] lemma symm_dual_id : bounded_order_hom.dual.symm (bounded_order_hom.id _) = bounded_order_hom.id α := rfl -@[simp] lemma symm_dual_comp (g : bounded_order_hom (order_dual β) (order_dual γ)) - (f : bounded_order_hom (order_dual α) (order_dual β)) : +@[simp] lemma symm_dual_comp (g : bounded_order_hom βᵒᵈ γᵒᵈ) (f : bounded_order_hom αᵒᵈ βᵒᵈ) : bounded_order_hom.dual.symm (g.comp f) = (bounded_order_hom.dual.symm g).comp (bounded_order_hom.dual.symm f) := rfl diff --git a/src/order/hom/complete_lattice.lean b/src/order/hom/complete_lattice.lean index ccec7e6f74eba..596d043ba064f 100644 --- a/src/order/hom/complete_lattice.lean +++ b/src/order/hom/complete_lattice.lean @@ -484,7 +484,7 @@ namespace Sup_hom variables [has_Sup α] [has_Sup β] [has_Sup γ] /-- Reinterpret a `⨆`-homomorphism as an `⨅`-homomorphism between the dual orders. -/ -@[simps] protected def dual : Sup_hom α β ≃ Inf_hom (order_dual α) (order_dual β) := +@[simps] protected def dual : Sup_hom α β ≃ Inf_hom αᵒᵈ βᵒᵈ := { to_fun := λ f, ⟨to_dual ∘ f ∘ of_dual, f.map_Sup'⟩, inv_fun := λ f, ⟨of_dual ∘ f ∘ to_dual, f.map_Inf'⟩, left_inv := λ f, Sup_hom.ext $ λ a, rfl, @@ -495,8 +495,7 @@ variables [has_Sup α] [has_Sup β] [has_Sup γ] (g.comp f).dual = g.dual.comp f.dual := rfl @[simp] lemma symm_dual_id : Sup_hom.dual.symm (Inf_hom.id _) = Sup_hom.id α := rfl -@[simp] lemma symm_dual_comp (g : Inf_hom (order_dual β) (order_dual γ)) - (f : Inf_hom (order_dual α) (order_dual β)) : +@[simp] lemma symm_dual_comp (g : Inf_hom βᵒᵈ γᵒᵈ) (f : Inf_hom αᵒᵈ βᵒᵈ) : Sup_hom.dual.symm (g.comp f) = (Sup_hom.dual.symm g).comp (Sup_hom.dual.symm f) := rfl end Sup_hom @@ -505,7 +504,7 @@ namespace Inf_hom variables [has_Inf α] [has_Inf β] [has_Inf γ] /-- Reinterpret an `⨅`-homomorphism as a `⨆`-homomorphism between the dual orders. -/ -@[simps] protected def dual : Inf_hom α β ≃ Sup_hom (order_dual α) (order_dual β) := +@[simps] protected def dual : Inf_hom α β ≃ Sup_hom αᵒᵈ βᵒᵈ := { to_fun := λ f, { to_fun := to_dual ∘ f ∘ of_dual, map_Sup' := λ _, congr_arg to_dual (map_Inf f _) }, inv_fun := λ f, { to_fun := of_dual ∘ f ∘ to_dual, @@ -518,8 +517,7 @@ variables [has_Inf α] [has_Inf β] [has_Inf γ] (g.comp f).dual = g.dual.comp f.dual := rfl @[simp] lemma symm_dual_id : Inf_hom.dual.symm (Sup_hom.id _) = Inf_hom.id α := rfl -@[simp] lemma symm_dual_comp (g : Sup_hom (order_dual β) (order_dual γ)) - (f : Sup_hom (order_dual α) (order_dual β)) : +@[simp] lemma symm_dual_comp (g : Sup_hom βᵒᵈ γᵒᵈ) (f : Sup_hom αᵒᵈ βᵒᵈ) : Inf_hom.dual.symm (g.comp f) = (Inf_hom.dual.symm g).comp (Inf_hom.dual.symm f) := rfl end Inf_hom @@ -529,8 +527,7 @@ variables [complete_lattice α] [complete_lattice β] [complete_lattice γ] /-- Reinterpret a complete lattice homomorphism as a complete lattice homomorphism between the dual lattices. -/ -@[simps] protected def dual : - complete_lattice_hom α β ≃ complete_lattice_hom (order_dual α) (order_dual β) := +@[simps] protected def dual : complete_lattice_hom α β ≃ complete_lattice_hom αᵒᵈ βᵒᵈ := { to_fun := λ f, ⟨f.to_Sup_hom.dual, f.map_Inf'⟩, inv_fun := λ f, ⟨f.to_Sup_hom.dual, f.map_Inf'⟩, left_inv := λ f, ext $ λ a, rfl, @@ -542,8 +539,7 @@ lattices. -/ @[simp] lemma symm_dual_id : complete_lattice_hom.dual.symm (complete_lattice_hom.id _) = complete_lattice_hom.id α := rfl -@[simp] lemma symm_dual_comp (g : complete_lattice_hom (order_dual β) (order_dual γ)) - (f : complete_lattice_hom (order_dual α) (order_dual β)) : +@[simp] lemma symm_dual_comp (g : complete_lattice_hom βᵒᵈ γᵒᵈ) (f : complete_lattice_hom αᵒᵈ βᵒᵈ) : complete_lattice_hom.dual.symm (g.comp f) = (complete_lattice_hom.dual.symm g).comp (complete_lattice_hom.dual.symm f) := rfl diff --git a/src/order/hom/lattice.lean b/src/order/hom/lattice.lean index 1721e0128c3c9..214fb509677d1 100644 --- a/src/order/hom/lattice.lean +++ b/src/order/hom/lattice.lean @@ -821,7 +821,7 @@ namespace sup_hom variables [has_sup α] [has_sup β] [has_sup γ] /-- Reinterpret a supremum homomorphism as an infimum homomorphism between the dual lattices. -/ -@[simps] protected def dual : sup_hom α β ≃ inf_hom (order_dual α) (order_dual β) := +@[simps] protected def dual : sup_hom α β ≃ inf_hom αᵒᵈ βᵒᵈ := { to_fun := λ f, ⟨f, f.map_sup'⟩, inv_fun := λ f, ⟨f, f.map_inf'⟩, left_inv := λ f, sup_hom.ext $ λ _, rfl, @@ -832,8 +832,7 @@ variables [has_sup α] [has_sup β] [has_sup γ] (g.comp f).dual = g.dual.comp f.dual := rfl @[simp] lemma symm_dual_id : sup_hom.dual.symm (inf_hom.id _) = sup_hom.id α := rfl -@[simp] lemma symm_dual_comp (g : inf_hom (order_dual β) (order_dual γ)) - (f : inf_hom (order_dual α) (order_dual β)) : +@[simp] lemma symm_dual_comp (g : inf_hom βᵒᵈ γᵒᵈ) (f : inf_hom αᵒᵈ βᵒᵈ) : sup_hom.dual.symm (g.comp f) = (sup_hom.dual.symm g).comp (sup_hom.dual.symm f) := rfl end sup_hom @@ -842,7 +841,7 @@ namespace inf_hom variables [has_inf α] [has_inf β] [has_inf γ] /-- Reinterpret an infimum homomorphism as a supremum homomorphism between the dual lattices. -/ -@[simps] protected def dual : inf_hom α β ≃ sup_hom (order_dual α) (order_dual β) := +@[simps] protected def dual : inf_hom α β ≃ sup_hom αᵒᵈ βᵒᵈ := { to_fun := λ f, ⟨f, f.map_inf'⟩, inv_fun := λ f, ⟨f, f.map_sup'⟩, left_inv := λ f, inf_hom.ext $ λ _, rfl, @@ -853,8 +852,7 @@ variables [has_inf α] [has_inf β] [has_inf γ] (g.comp f).dual = g.dual.comp f.dual := rfl @[simp] lemma symm_dual_id : inf_hom.dual.symm (sup_hom.id _) = inf_hom.id α := rfl -@[simp] lemma symm_dual_comp (g : sup_hom (order_dual β) (order_dual γ)) - (f : sup_hom (order_dual α) (order_dual β)) : +@[simp] lemma symm_dual_comp (g : sup_hom βᵒᵈ γᵒᵈ) (f : sup_hom αᵒᵈ βᵒᵈ) : inf_hom.dual.symm (g.comp f) = (inf_hom.dual.symm g).comp (inf_hom.dual.symm f) := rfl end inf_hom @@ -864,7 +862,7 @@ variables [has_sup α] [has_bot α] [has_sup β] [has_bot β] [has_sup γ] [has_ /-- Reinterpret a finitary supremum homomorphism as a finitary infimum homomorphism between the dual lattices. -/ -def dual : sup_bot_hom α β ≃ inf_top_hom (order_dual α) (order_dual β) := +def dual : sup_bot_hom α β ≃ inf_top_hom αᵒᵈ βᵒᵈ := { to_fun := λ f, ⟨f.to_sup_hom.dual, f.map_bot'⟩, inv_fun := λ f, ⟨sup_hom.dual.symm f.to_inf_hom, f.map_top'⟩, left_inv := λ f, sup_bot_hom.ext $ λ _, rfl, @@ -875,8 +873,7 @@ def dual : sup_bot_hom α β ≃ inf_top_hom (order_dual α) (order_dual β) := (g.comp f).dual = g.dual.comp f.dual := rfl @[simp] lemma symm_dual_id : sup_bot_hom.dual.symm (inf_top_hom.id _) = sup_bot_hom.id α := rfl -@[simp] lemma symm_dual_comp (g : inf_top_hom (order_dual β) (order_dual γ)) - (f : inf_top_hom (order_dual α) (order_dual β)) : +@[simp] lemma symm_dual_comp (g : inf_top_hom βᵒᵈ γᵒᵈ) (f : inf_top_hom αᵒᵈ βᵒᵈ) : sup_bot_hom.dual.symm (g.comp f) = (sup_bot_hom.dual.symm g).comp (sup_bot_hom.dual.symm f) := rfl end sup_bot_hom @@ -886,7 +883,7 @@ variables [has_inf α] [has_top α] [has_inf β] [has_top β] [has_inf γ] [has_ /-- Reinterpret a finitary infimum homomorphism as a finitary supremum homomorphism between the dual lattices. -/ -@[simps] protected def dual : inf_top_hom α β ≃ sup_bot_hom (order_dual α) (order_dual β) := +@[simps] protected def dual : inf_top_hom α β ≃ sup_bot_hom αᵒᵈ βᵒᵈ := { to_fun := λ f, ⟨f.to_inf_hom.dual, f.map_top'⟩, inv_fun := λ f, ⟨inf_hom.dual.symm f.to_sup_hom, f.map_bot'⟩, left_inv := λ f, inf_top_hom.ext $ λ _, rfl, @@ -897,8 +894,7 @@ lattices. -/ (g.comp f).dual = g.dual.comp f.dual := rfl @[simp] lemma symm_dual_id : inf_top_hom.dual.symm (sup_bot_hom.id _) = inf_top_hom.id α := rfl -@[simp] lemma symm_dual_comp (g : sup_bot_hom (order_dual β) (order_dual γ)) - (f : sup_bot_hom (order_dual α) (order_dual β)) : +@[simp] lemma symm_dual_comp (g : sup_bot_hom βᵒᵈ γᵒᵈ) (f : sup_bot_hom αᵒᵈ βᵒᵈ) : inf_top_hom.dual.symm (g.comp f) = (inf_top_hom.dual.symm g).comp (inf_top_hom.dual.symm f) := rfl end inf_top_hom @@ -907,7 +903,7 @@ namespace lattice_hom variables [lattice α] [lattice β] [lattice γ] /-- Reinterpret a lattice homomorphism as a lattice homomorphism between the dual lattices. -/ -@[simps] protected def dual : lattice_hom α β ≃ lattice_hom (order_dual α) (order_dual β) := +@[simps] protected def dual : lattice_hom α β ≃ lattice_hom αᵒᵈ βᵒᵈ := { to_fun := λ f, ⟨f.to_inf_hom.dual, f.map_sup'⟩, inv_fun := λ f, ⟨f.to_inf_hom.dual, f.map_sup'⟩, left_inv := λ f, ext $ λ a, rfl, @@ -918,8 +914,7 @@ variables [lattice α] [lattice β] [lattice γ] (g.comp f).dual = g.dual.comp f.dual := rfl @[simp] lemma symm_dual_id : lattice_hom.dual.symm (lattice_hom.id _) = lattice_hom.id α := rfl -@[simp] lemma symm_dual_comp (g : lattice_hom (order_dual β) (order_dual γ)) - (f : lattice_hom (order_dual α) (order_dual β)) : +@[simp] lemma symm_dual_comp (g : lattice_hom βᵒᵈ γᵒᵈ) (f : lattice_hom αᵒᵈ βᵒᵈ) : lattice_hom.dual.symm (g.comp f) = (lattice_hom.dual.symm g).comp (lattice_hom.dual.symm f) := rfl end lattice_hom @@ -929,8 +924,7 @@ variables [lattice α] [bounded_order α] [lattice β] [bounded_order β] [latti /-- Reinterpret a bounded lattice homomorphism as a bounded lattice homomorphism between the dual bounded lattices. -/ -@[simps] protected def dual : - bounded_lattice_hom α β ≃ bounded_lattice_hom (order_dual α) (order_dual β) := +@[simps] protected def dual : bounded_lattice_hom α β ≃ bounded_lattice_hom αᵒᵈ βᵒᵈ := { to_fun := λ f, ⟨f.to_lattice_hom.dual, f.map_bot', f.map_top'⟩, inv_fun := λ f, ⟨lattice_hom.dual.symm f.to_lattice_hom, f.map_bot', f.map_top'⟩, left_inv := λ f, ext $ λ a, rfl, @@ -942,8 +936,7 @@ bounded lattices. -/ @[simp] lemma symm_dual_id : bounded_lattice_hom.dual.symm (bounded_lattice_hom.id _) = bounded_lattice_hom.id α := rfl -@[simp] lemma symm_dual_comp (g : bounded_lattice_hom (order_dual β) (order_dual γ)) - (f : bounded_lattice_hom (order_dual α) (order_dual β)) : +@[simp] lemma symm_dual_comp (g : bounded_lattice_hom βᵒᵈ γᵒᵈ) (f : bounded_lattice_hom αᵒᵈ βᵒᵈ) : bounded_lattice_hom.dual.symm (g.comp f) = (bounded_lattice_hom.dual.symm g).comp (bounded_lattice_hom.dual.symm f) := rfl diff --git a/src/order/iterate.lean b/src/order/iterate.lean index bcd0164440a1a..2b0b78c82ce61 100644 --- a/src/order/iterate.lean +++ b/src/order/iterate.lean @@ -123,13 +123,13 @@ lemma id_le_iterate_of_id_le (h : id ≤ f) (n : ℕ) : id ≤ (f^[n]) := by simpa only [iterate_id] using monotone_id.iterate_le_of_le h n lemma iterate_le_id_of_le_id (h : f ≤ id) (n : ℕ) : (f^[n]) ≤ id := -@id_le_iterate_of_id_le (order_dual α) _ f h n +@id_le_iterate_of_id_le αᵒᵈ _ f h n lemma monotone_iterate_of_id_le (h : id ≤ f) : monotone (λ m, f^[m]) := monotone_nat_of_le_succ $ λ n x, by { rw iterate_succ_apply', exact h _ } lemma antitone_iterate_of_le_id (h : f ≤ id) : antitone (λ m, f^[m]) := -λ m n hmn, @monotone_iterate_of_id_le (order_dual α) _ f h m n hmn +λ m n hmn, @monotone_iterate_of_id_le αᵒᵈ _ f h m n hmn end preorder @@ -161,7 +161,7 @@ by refine hf.seq_pos_lt_seq_of_le_of_lt hn _ (λ k hk, _) (λ k hk, _); lemma iterate_pos_lt_of_map_lt' (h : commute f g) (hf : strict_mono f) (hg : monotone g) {x} (hx : f x < g x) {n} (hn : 0 < n) : f^[n] x < (g^[n]) x := -@iterate_pos_lt_of_map_lt (order_dual α) _ g f h.symm hg.dual hf.dual x hx n hn +@iterate_pos_lt_of_map_lt αᵒᵈ _ g f h.symm hg.dual hf.dual x hx n hn end preorder @@ -180,7 +180,7 @@ end lemma iterate_pos_lt_iff_map_lt' (h : commute f g) (hf : strict_mono f) (hg : monotone g) {x n} (hn : 0 < n) : f^[n] x < (g^[n]) x ↔ f x < g x := -@iterate_pos_lt_iff_map_lt (order_dual α) _ _ _ h.symm hg.dual hf.dual x n hn +@iterate_pos_lt_iff_map_lt αᵒᵈ _ _ _ h.symm hg.dual hf.dual x n hn lemma iterate_pos_le_iff_map_le (h : commute f g) (hf : monotone f) (hg : strict_mono g) {x n} (hn : 0 < n) : diff --git a/src/order/lattice.lean b/src/order/lattice.lean index a0aeec33efa9e..c59a20e18ac9a 100644 --- a/src/order/lattice.lean +++ b/src/order/lattice.lean @@ -119,8 +119,8 @@ def semilattice_sup.mk' {α : Type*} [has_sup α] rwa [sup_assoc, hbc], end } -instance (α : Type*) [has_inf α] : has_sup (order_dual α) := ⟨((⊓) : α → α → α)⟩ -instance (α : Type*) [has_sup α] : has_inf (order_dual α) := ⟨((⊔) : α → α → α)⟩ +instance (α : Type*) [has_inf α] : has_sup αᵒᵈ := ⟨((⊓) : α → α → α)⟩ +instance (α : Type*) [has_sup α] : has_inf αᵒᵈ := ⟨((⊔) : α → α → α)⟩ section semilattice_sup variables [semilattice_sup α] {a b c d : α} @@ -284,20 +284,20 @@ class semilattice_inf (α : Type u) extends has_inf α, partial_order α := (inf_le_right : ∀ a b : α, a ⊓ b ≤ b) (le_inf : ∀ a b c : α, a ≤ b → a ≤ c → a ≤ b ⊓ c) -instance (α) [semilattice_inf α] : semilattice_sup (order_dual α) := +instance (α) [semilattice_inf α] : semilattice_sup αᵒᵈ := { le_sup_left := semilattice_inf.inf_le_left, le_sup_right := semilattice_inf.inf_le_right, sup_le := assume a b c hca hcb, @semilattice_inf.le_inf α _ _ _ _ hca hcb, .. order_dual.partial_order α, .. order_dual.has_sup α } -instance (α) [semilattice_sup α] : semilattice_inf (order_dual α) := +instance (α) [semilattice_sup α] : semilattice_inf αᵒᵈ := { inf_le_left := @le_sup_left α _, inf_le_right := @le_sup_right α _, le_inf := assume a b c hca hcb, @sup_le α _ _ _ _ hca hcb, .. order_dual.partial_order α, .. order_dual.has_inf α } theorem semilattice_sup.dual_dual (α : Type*) [H : semilattice_sup α] : - order_dual.semilattice_sup (order_dual α) = H := + order_dual.semilattice_sup αᵒᵈ = H := semilattice_sup.ext $ λ _ _, iff.rfl section semilattice_inf @@ -330,8 +330,7 @@ lt_of_le_of_lt inf_le_left h theorem inf_lt_of_right_lt (h : b < c) : a ⊓ b < c := lt_of_le_of_lt inf_le_right h -@[simp] theorem le_inf_iff : a ≤ b ⊓ c ↔ a ≤ b ∧ a ≤ c := -@sup_le_iff (order_dual α) _ _ _ _ +@[simp] theorem le_inf_iff : a ≤ b ⊓ c ↔ a ≤ b ∧ a ≤ c := @sup_le_iff αᵒᵈ _ _ _ _ @[simp] theorem inf_eq_left : a ⊓ b = a ↔ a ≤ b := le_antisymm_iff.trans $ by simp [le_refl] @@ -363,18 +362,15 @@ inf_le_inf le_rfl h theorem le_of_inf_eq (h : a ⊓ b = a) : a ≤ b := by { rw ← h, simp } -@[simp] theorem inf_idem : a ⊓ a = a := -@sup_idem (order_dual α) _ _ +@[simp] lemma inf_idem : a ⊓ a = a := @sup_idem αᵒᵈ _ _ instance inf_is_idempotent : is_idempotent α (⊓) := ⟨@inf_idem _ _⟩ -theorem inf_comm : a ⊓ b = b ⊓ a := -@sup_comm (order_dual α) _ _ _ +lemma inf_comm : a ⊓ b = b ⊓ a := @sup_comm αᵒᵈ _ _ _ instance inf_is_commutative : is_commutative α (⊓) := ⟨@inf_comm _ _⟩ -theorem inf_assoc : a ⊓ b ⊓ c = a ⊓ (b ⊓ c) := -@sup_assoc (order_dual α) _ a b c +lemma inf_assoc : a ⊓ b ⊓ c = a ⊓ (b ⊓ c) := @sup_assoc αᵒᵈ _ a b c instance inf_is_associative : is_associative α (⊓) := ⟨@inf_assoc _ _⟩ @@ -382,28 +378,28 @@ lemma inf_left_right_swap (a b c : α) : a ⊓ b ⊓ c = c ⊓ b ⊓ a := by rw [inf_comm, @inf_comm _ _ a, inf_assoc] @[simp] lemma inf_left_idem : a ⊓ (a ⊓ b) = a ⊓ b := -@sup_left_idem (order_dual α) _ a b +@sup_left_idem αᵒᵈ _ a b @[simp] lemma inf_right_idem : (a ⊓ b) ⊓ b = a ⊓ b := -@sup_right_idem (order_dual α) _ a b +@sup_right_idem αᵒᵈ _ a b lemma inf_left_comm (a b c : α) : a ⊓ (b ⊓ c) = b ⊓ (a ⊓ c) := -@sup_left_comm (order_dual α) _ a b c +@sup_left_comm αᵒᵈ _ a b c lemma inf_right_comm (a b c : α) : a ⊓ b ⊓ c = a ⊓ c ⊓ b := -@sup_right_comm (order_dual α) _ a b c +@sup_right_comm αᵒᵈ _ a b c lemma inf_inf_inf_comm (a b c d : α) : a ⊓ b ⊓ (c ⊓ d) = a ⊓ c ⊓ (b ⊓ d) := -@sup_sup_sup_comm (order_dual α) _ _ _ _ _ +@sup_sup_sup_comm αᵒᵈ _ _ _ _ _ lemma inf_inf_distrib_left (a b c : α) : a ⊓ (b ⊓ c) = (a ⊓ b) ⊓ (a ⊓ c) := -@sup_sup_distrib_left (order_dual α) _ _ _ _ +@sup_sup_distrib_left αᵒᵈ _ _ _ _ lemma inf_inf_distrib_right (a b c : α) : (a ⊓ b) ⊓ c = (a ⊓ c) ⊓ (b ⊓ c) := -@sup_sup_distrib_right (order_dual α) _ _ _ _ +@sup_sup_distrib_right αᵒᵈ _ _ _ _ lemma forall_le_or_exists_lt_inf (a : α) : (∀b, a ≤ b) ∨ (∃b, b < a) := -@forall_le_or_exists_lt_sup (order_dual α) _ a +@forall_le_or_exists_lt_sup αᵒᵈ _ a theorem semilattice_inf.ext_inf {α} {A B : semilattice_inf α} (H : ∀ x y : α, (by haveI := A; exact x ≤ y) ↔ x ≤ y) @@ -421,12 +417,12 @@ begin end theorem semilattice_inf.dual_dual (α : Type*) [H : semilattice_inf α] : - order_dual.semilattice_inf (order_dual α) = H := + order_dual.semilattice_inf αᵒᵈ = H := semilattice_inf.ext $ λ _ _, iff.rfl theorem exists_lt_of_inf (α : Type*) [semilattice_inf α] [nontrivial α] : ∃ a b : α, a < b := -let ⟨a, b, h⟩ := exists_lt_of_sup (order_dual α) in ⟨b, a, h⟩ +let ⟨a, b, h⟩ := exists_lt_of_sup αᵒᵈ in ⟨b, a, h⟩ lemma inf_le_ite (s s' : α) (P : Prop) [decidable P] : s ⊓ s' ≤ ite P s s' := if h : P then inf_le_left.trans_eq (if_pos h).symm else inf_le_right.trans_eq (if_neg h).symm @@ -444,8 +440,8 @@ def semilattice_inf.mk' {α : Type*} [has_inf α] (inf_assoc : ∀ (a b c : α), a ⊓ b ⊓ c = a ⊓ (b ⊓ c)) (inf_idem : ∀ (a : α), a ⊓ a = a) : semilattice_inf α := begin - haveI : semilattice_sup (order_dual α) := semilattice_sup.mk' inf_comm inf_assoc inf_idem, - haveI i := order_dual.semilattice_inf (order_dual α), + haveI : semilattice_sup αᵒᵈ := semilattice_sup.mk' inf_comm inf_assoc inf_idem, + haveI i := order_dual.semilattice_inf αᵒᵈ, exact i, end @@ -456,7 +452,7 @@ end /-- A lattice is a join-semilattice which is also a meet-semilattice. -/ @[protect_proj] class lattice (α : Type u) extends semilattice_sup α, semilattice_inf α -instance (α) [lattice α] : lattice (order_dual α) := +instance (α) [lattice α] : lattice αᵒᵈ := { .. order_dual.semilattice_sup α, .. order_dual.semilattice_inf α } /-- The partial orders from `semilattice_sup_mk'` and `semilattice_inf_mk'` agree @@ -598,7 +594,7 @@ calc x ⊓ (y ⊔ z) = (x ⊓ (x ⊔ z)) ⊓ (y ⊔ z) : by rw [inf_sup_se ... = ((x ⊓ y) ⊔ x) ⊓ ((x ⊓ y) ⊔ z) : by rw [sup_comm] ... = (x ⊓ y) ⊔ (x ⊓ z) : by rw [sup_inf_left] -instance (α : Type*) [distrib_lattice α] : distrib_lattice (order_dual α) := +instance (α : Type*) [distrib_lattice α] : distrib_lattice αᵒᵈ := { le_sup_inf := assume x y z, le_of_eq inf_sup_left.symm, .. order_dual.lattice α } @@ -663,12 +659,11 @@ lemma sup_ind (a b : α) {p : α → Prop} (ha : p a) (hb : p b) : p (a ⊔ b) : @[simp] lemma sup_lt_iff : b ⊔ c < a ↔ b < a ∧ c < a := ⟨λ h, ⟨le_sup_left.trans_lt h, le_sup_right.trans_lt h⟩, λ h, sup_ind b c h.1 h.2⟩ -lemma inf_ind (a b : α) {p : α → Prop} (ha : p a) (hb : p b) : p (a ⊓ b) := -@sup_ind (order_dual α) _ _ _ _ ha hb +lemma inf_ind (a b : α) {p : α → Prop} : p a → p b → p (a ⊓ b) := @sup_ind αᵒᵈ _ _ _ _ -@[simp] lemma inf_le_iff : b ⊓ c ≤ a ↔ b ≤ a ∨ c ≤ a := @le_sup_iff (order_dual α) _ _ _ _ -@[simp] lemma inf_lt_iff : b ⊓ c < a ↔ b < a ∨ c < a := @lt_sup_iff (order_dual α) _ _ _ _ -@[simp] lemma lt_inf_iff : a < b ⊓ c ↔ a < b ∧ a < c := @sup_lt_iff (order_dual α) _ _ _ _ +@[simp] lemma inf_le_iff : b ⊓ c ≤ a ↔ b ≤ a ∨ c ≤ a := @le_sup_iff αᵒᵈ _ _ _ _ +@[simp] lemma inf_lt_iff : b ⊓ c < a ↔ b < a ∨ c < a := @lt_sup_iff αᵒᵈ _ _ _ _ +@[simp] lemma lt_inf_iff : a < b ⊓ c ↔ a < b ∧ a < c := @sup_lt_iff αᵒᵈ _ _ _ _ end linear_order @@ -683,7 +678,7 @@ end lemma inf_eq_min_default [semilattice_inf α] [decidable_rel ((≤) : α → α → Prop)] [is_total α (≤)] : (⊓) = (min_default : α → α → α) := -@sup_eq_max_default (order_dual α) _ _ _ +@sup_eq_max_default αᵒᵈ _ _ _ /-- A lattice with total order is a linear order. diff --git a/src/order/liminf_limsup.lean b/src/order/liminf_limsup.lean index 19fd5d35a6d4c..f80f61d5317e2 100644 --- a/src/order/liminf_limsup.lean +++ b/src/order/liminf_limsup.lean @@ -105,7 +105,7 @@ end lemma not_is_bounded_under_of_tendsto_at_bot [preorder β] [no_min_order β] {f : α → β} {l : filter α} [l.ne_bot](hf : tendsto f l at_bot) : ¬ is_bounded_under (≥) l f := -@not_is_bounded_under_of_tendsto_at_top α (order_dual β) _ _ _ _ _ hf +@not_is_bounded_under_of_tendsto_at_top α βᵒᵈ _ _ _ _ _ hf lemma is_bounded_under.bdd_above_range_of_cofinite [semilattice_sup β] {f : α → β} (hf : is_bounded_under (≤) cofinite f) : bdd_above (range f) := @@ -118,7 +118,7 @@ end lemma is_bounded_under.bdd_below_range_of_cofinite [semilattice_inf β] {f : α → β} (hf : is_bounded_under (≥) cofinite f) : bdd_below (range f) := -@is_bounded_under.bdd_above_range_of_cofinite α (order_dual β) _ _ hf +@is_bounded_under.bdd_above_range_of_cofinite α βᵒᵈ _ _ hf lemma is_bounded_under.bdd_above_range [semilattice_sup β] {f : ℕ → β} (hf : is_bounded_under (≤) at_top f) : bdd_above (range f) := @@ -126,7 +126,7 @@ by { rw ← nat.cofinite_eq_at_top at hf, exact hf.bdd_above_range_of_cofinite } lemma is_bounded_under.bdd_below_range [semilattice_inf β] {f : ℕ → β} (hf : is_bounded_under (≥) at_top f) : bdd_below (range f) := -@is_bounded_under.bdd_above_range (order_dual β) _ _ hf +@is_bounded_under.bdd_above_range βᵒᵈ _ _ hf /-- `is_cobounded (≺) f` states that the filter `f` does not tend to infinity w.r.t. `≺`. This is also called frequently bounded. Will be usually instantiated with `≤` or `≥`. @@ -299,7 +299,7 @@ lemma liminf_le_liminf {α : Type*} [conditionally_complete_lattice β] {f : fil (hu : f.is_bounded_under (≥) u . is_bounded_default) (hv : f.is_cobounded_under (≥) v . is_bounded_default) : f.liminf u ≤ f.liminf v := -@limsup_le_limsup (order_dual β) α _ _ _ _ h hv hu +@limsup_le_limsup βᵒᵈ α _ _ _ _ h hv hu lemma limsup_le_limsup_of_le {α β} [conditionally_complete_lattice β] {f g : filter α} (h : f ≤ g) {u : α → β} (hf : f.is_cobounded_under (≤) u . is_bounded_default) @@ -319,7 +319,7 @@ by simp [Limsup]; exact cInf_upper_bounds_eq_cSup h hs theorem Liminf_principal {s : set α} (h : bdd_below s) (hs : s.nonempty) : (𝓟 s).Liminf = Inf s := -@Limsup_principal (order_dual α) _ s h hs +@Limsup_principal αᵒᵈ _ s h hs lemma limsup_congr {α : Type*} [conditionally_complete_lattice β] {f : filter α} {u v : α → β} (h : ∀ᶠ a in f, u a = v a) : limsup f u = limsup f v := @@ -331,7 +331,7 @@ end lemma liminf_congr {α : Type*} [conditionally_complete_lattice β] {f : filter α} {u v : α → β} (h : ∀ᶠ a in f, u a = v a) : liminf f u = liminf f v := -@limsup_congr (order_dual β) _ _ _ _ _ h +@limsup_congr βᵒᵈ _ _ _ _ _ h lemma limsup_const {α : Type*} [conditionally_complete_lattice β] {f : filter α} [ne_bot f] (b : β) : limsup f (λ x, b) = b := @@ -339,7 +339,7 @@ by simpa only [limsup_eq, eventually_const] using cInf_Ici lemma liminf_const {α : Type*} [conditionally_complete_lattice β] {f : filter α} [ne_bot f] (b : β) : liminf f (λ x, b) = b := -@limsup_const (order_dual β) α _ f _ b +@limsup_const βᵒᵈ α _ f _ b lemma liminf_le_limsup {f : filter β} [ne_bot f] {u : β → α} (h : f.is_bounded_under (≤) u . is_bounded_default) @@ -375,7 +375,7 @@ end /-- Same as limsup_const applied to `⊤` but without the `ne_bot f` assumption -/ lemma liminf_const_top {f : filter β} : liminf f (λ x : β, (⊤ : α)) = (⊤ : α) := -@limsup_const_bot (order_dual α) β _ _ +@limsup_const_bot αᵒᵈ β _ _ theorem has_basis.Limsup_eq_infi_Sup {ι} {p : ι → Prop} {s} {f : filter α} (h : f.has_basis p s) : f.Limsup = ⨅ i (hi : p i), Sup (s i) := @@ -386,13 +386,13 @@ le_antisymm theorem has_basis.Liminf_eq_supr_Inf {p : ι → Prop} {s : ι → set α} {f : filter α} (h : f.has_basis p s) : f.Liminf = ⨆ i (hi : p i), Inf (s i) := -@has_basis.Limsup_eq_infi_Sup (order_dual α) _ _ _ _ _ h +@has_basis.Limsup_eq_infi_Sup αᵒᵈ _ _ _ _ _ h theorem Limsup_eq_infi_Sup {f : filter α} : f.Limsup = ⨅ s ∈ f, Sup s := f.basis_sets.Limsup_eq_infi_Sup theorem Liminf_eq_supr_Inf {f : filter α} : f.Liminf = ⨆ s ∈ f, Inf s := -@Limsup_eq_infi_Sup (order_dual α) _ _ +@Limsup_eq_infi_Sup αᵒᵈ _ _ /-- In a complete lattice, the limsup of a function is the infimum over sets `s` in the filter of the supremum of the function over `s` -/ @@ -414,17 +414,17 @@ theorem has_basis.limsup_eq_infi_supr {p : ι → Prop} {s : ι → set β} {f : /-- In a complete lattice, the liminf of a function is the infimum over sets `s` in the filter of the supremum of the function over `s` -/ theorem liminf_eq_supr_infi {f : filter β} {u : β → α} : f.liminf u = ⨆ s ∈ f, ⨅ a ∈ s, u a := -@limsup_eq_infi_supr (order_dual α) β _ _ _ +@limsup_eq_infi_supr αᵒᵈ β _ _ _ lemma liminf_eq_supr_infi_of_nat {u : ℕ → α} : liminf at_top u = ⨆ n : ℕ, ⨅ i ≥ n, u i := -@limsup_eq_infi_supr_of_nat (order_dual α) _ u +@limsup_eq_infi_supr_of_nat αᵒᵈ _ u lemma liminf_eq_supr_infi_of_nat' {u : ℕ → α} : liminf at_top u = ⨆ n : ℕ, ⨅ i : ℕ, u (i + n) := -@limsup_eq_infi_supr_of_nat' (order_dual α) _ _ +@limsup_eq_infi_supr_of_nat' αᵒᵈ _ _ theorem has_basis.liminf_eq_supr_infi {p : ι → Prop} {s : ι → set β} {f : filter β} {u : β → α} (h : f.has_basis p s) : f.liminf u = ⨆ i (hi : p i), ⨅ a ∈ s i, u a := -@has_basis.limsup_eq_infi_supr (order_dual α) _ _ _ _ _ _ _ h +@has_basis.limsup_eq_infi_supr αᵒᵈ _ _ _ _ _ _ _ h @[simp] lemma liminf_nat_add (f : ℕ → α) (k : ℕ) : at_top.liminf (λ i, f (i + k)) = at_top.liminf f := @@ -432,7 +432,7 @@ by { simp_rw liminf_eq_supr_infi_of_nat, exact supr_infi_ge_nat_add f k } @[simp] lemma limsup_nat_add (f : ℕ → α) (k : ℕ) : at_top.limsup (λ i, f (i + k)) = at_top.limsup f := -@liminf_nat_add (order_dual α) _ f k +@liminf_nat_add αᵒᵈ _ f k lemma liminf_le_of_frequently_le' {α β} [complete_lattice β] {f : filter α} {u : α → β} {x : β} (h : ∃ᶠ a in f, u a ≤ x) : @@ -450,7 +450,7 @@ end lemma le_limsup_of_frequently_le' {α β} [complete_lattice β] {f : filter α} {u : α → β} {x : β} (h : ∃ᶠ a in f, x ≤ u a) : x ≤ f.limsup u := -@liminf_le_of_frequently_le' _ (order_dual β) _ _ _ _ h +@liminf_le_of_frequently_le' _ βᵒᵈ _ _ _ _ h end complete_lattice @@ -468,7 +468,7 @@ end lemma eventually_lt_of_limsup_lt {f : filter α} [conditionally_complete_linear_order β] {u : α → β} {b : β} (h : limsup f u < b) (hu : f.is_bounded_under (≤) u . is_bounded_default) : ∀ᶠ a in f, u a < b := -@eventually_lt_of_lt_liminf _ (order_dual β) _ _ _ _ h hu +@eventually_lt_of_lt_liminf _ βᵒᵈ _ _ _ _ h hu lemma le_limsup_of_frequently_le {α β} [conditionally_complete_linear_order β] {f : filter α} {u : α → β} {b : β} (hu_le : ∃ᶠ x in f, b ≤ u x) @@ -485,7 +485,7 @@ lemma liminf_le_of_frequently_le {α β} [conditionally_complete_linear_order {u : α → β} {b : β} (hu_le : ∃ᶠ x in f, u x ≤ b) (hu : f.is_bounded_under (≥) u . is_bounded_default) : f.liminf u ≤ b := -@le_limsup_of_frequently_le _ (order_dual β) _ f u b hu_le hu +@le_limsup_of_frequently_le _ βᵒᵈ _ f u b hu_le hu lemma frequently_lt_of_lt_limsup {α β} [conditionally_complete_linear_order β] {f : filter α} {u : α → β} {b : β} @@ -501,7 +501,7 @@ lemma frequently_lt_of_liminf_lt {α β} [conditionally_complete_linear_order β {u : α → β} {b : β} (hu : f.is_cobounded_under (≥) u . is_bounded_default) (h : f.liminf u < b) : ∃ᶠ x in f, u x < b := -@frequently_lt_of_lt_limsup _ (order_dual β) _ f u b hu h +@frequently_lt_of_lt_limsup _ βᵒᵈ _ f u b hu h end conditionally_complete_linear_order @@ -548,6 +548,6 @@ lemma order_iso.liminf_apply {γ} [conditionally_complete_lattice β] (hgu : f.is_bounded_under (≥) (λ x, g (u x)) . is_bounded_default) (hgu_co : f.is_cobounded_under (≥) (λ x, g (u x)) . is_bounded_default) : g (f.liminf u) = f.liminf (λ x, g (u x)) := -@order_iso.limsup_apply α (order_dual β) (order_dual γ) _ _ f u g.dual hu hu_co hgu hgu_co +@order_iso.limsup_apply α βᵒᵈ γᵒᵈ _ _ f u g.dual hu hu_co hgu hgu_co end order diff --git a/src/order/locally_finite.lean b/src/order/locally_finite.lean index 937ba1a08cef6..54634155c3da1 100644 --- a/src/order/locally_finite.lean +++ b/src/order/locally_finite.lean @@ -423,13 +423,13 @@ open order_dual variables [locally_finite_order α] (a b : α) /-- Note we define `Icc (to_dual a) (to_dual b)` as `Icc α _ _ b a` (which has type `finset α` not -`finset (order_dual α)`!) instead of `(Icc b a).map to_dual.to_embedding` as this means the +`finset αᵒᵈ`!) instead of `(Icc b a).map to_dual.to_embedding` as this means the following is defeq: ``` lemma this : (Icc (to_dual (to_dual a)) (to_dual (to_dual b)) : _) = (Icc a b : _) := rfl ``` -/ -instance : locally_finite_order (order_dual α) := +instance : locally_finite_order αᵒᵈ := { finset_Icc := λ a b, @Icc α _ _ (of_dual b) (of_dual a), finset_Ico := λ a b, @Ioc α _ _ (of_dual b) (of_dual a), finset_Ioc := λ a b, @Ico α _ _ (of_dual b) (of_dual a), @@ -570,7 +570,7 @@ namespace with_bot variables (α) [partial_order α] [order_bot α] [locally_finite_order α] instance : locally_finite_order (with_bot α) := -@order_dual.locally_finite_order (with_top (order_dual α)) _ _ +@order_dual.locally_finite_order (with_top αᵒᵈ) _ _ variables (a b : α) diff --git a/src/order/max.lean b/src/order/max.lean index 56c53cab2685a..38f2de28e046f 100644 --- a/src/order/max.lean +++ b/src/order/max.lean @@ -59,20 +59,16 @@ nonempty_subtype.2 (exists_lt a) instance nonempty_gt [has_lt α] [no_max_order α] (a : α) : nonempty {x // a < x} := nonempty_subtype.2 (exists_gt a) -instance order_dual.no_bot_order (α : Type*) [has_le α] [no_top_order α] : - no_bot_order (order_dual α) := +instance order_dual.no_bot_order (α : Type*) [has_le α] [no_top_order α] : no_bot_order αᵒᵈ := ⟨λ a, @exists_not_le α _ _ a⟩ -instance order_dual.no_top_order (α : Type*) [has_le α] [no_bot_order α] : - no_top_order (order_dual α) := +instance order_dual.no_top_order (α : Type*) [has_le α] [no_bot_order α] : no_top_order αᵒᵈ := ⟨λ a, @exists_not_ge α _ _ a⟩ -instance order_dual.no_min_order (α : Type*) [has_lt α] [no_max_order α] : - no_min_order (order_dual α) := +instance order_dual.no_min_order (α : Type*) [has_lt α] [no_max_order α] : no_min_order αᵒᵈ := ⟨λ a, @exists_gt α _ _ a⟩ -instance order_dual.no_max_order (α : Type*) [has_lt α] [no_min_order α] : - no_max_order (order_dual α) := +instance order_dual.no_max_order (α : Type*) [has_lt α] [no_min_order α] : no_max_order αᵒᵈ := ⟨λ a, @exists_lt α _ _ a⟩ @[priority 100] -- See note [lower instance priority] @@ -121,10 +117,10 @@ protected lemma is_top.is_max (h : is_top a) : is_max a := λ b _, h b @[simp] lemma is_top_to_dual_iff : is_top (to_dual a) ↔ is_bot a := iff.rfl @[simp] lemma is_min_to_dual_iff : is_min (to_dual a) ↔ is_max a := iff.rfl @[simp] lemma is_max_to_dual_iff : is_max (to_dual a) ↔ is_min a := iff.rfl -@[simp] lemma is_bot_of_dual_iff {a : order_dual α} : is_bot (of_dual a) ↔ is_top a := iff.rfl -@[simp] lemma is_top_of_dual_iff {a : order_dual α} : is_top (of_dual a) ↔ is_bot a := iff.rfl -@[simp] lemma is_min_of_dual_iff {a : order_dual α} : is_min (of_dual a) ↔ is_max a := iff.rfl -@[simp] lemma is_max_of_dual_iff {a : order_dual α} : is_max (of_dual a) ↔ is_min a := iff.rfl +@[simp] lemma is_bot_of_dual_iff {a : αᵒᵈ} : is_bot (of_dual a) ↔ is_top a := iff.rfl +@[simp] lemma is_top_of_dual_iff {a : αᵒᵈ} : is_top (of_dual a) ↔ is_bot a := iff.rfl +@[simp] lemma is_min_of_dual_iff {a : αᵒᵈ} : is_min (of_dual a) ↔ is_max a := iff.rfl +@[simp] lemma is_max_of_dual_iff {a : αᵒᵈ} : is_max (of_dual a) ↔ is_min a := iff.rfl alias is_bot_to_dual_iff ↔ _ is_top.to_dual alias is_top_to_dual_iff ↔ _ is_bot.to_dual @@ -196,6 +192,6 @@ variables [linear_order α] lemma is_top_or_exists_gt (a : α) : is_top a ∨ ∃ b, a < b := by simpa only [or_iff_not_imp_left, is_top, not_forall, not_le] using id -lemma is_bot_or_exists_lt (a : α) : is_bot a ∨ ∃ b, b < a := @is_top_or_exists_gt (order_dual α) _ a +lemma is_bot_or_exists_lt (a : α) : is_bot a ∨ ∃ b, b < a := @is_top_or_exists_gt αᵒᵈ _ a end linear_order diff --git a/src/order/min_max.lean b/src/order/min_max.lean index cbe72653f0ef4..46d96d9a40b4d 100644 --- a/src/order/min_max.lean +++ b/src/order/min_max.lean @@ -68,8 +68,7 @@ end /-- For elements `a` and `b` of a linear order, either `max a b = a` and `b ≤ a`, or `max a b = b` and `a < b`. Use cases on this lemma to automate linarith in inequalities -/ -lemma max_cases (a b : α) : max a b = a ∧ b ≤ a ∨ max a b = b ∧ a < b := -@min_cases (order_dual α) _ a b +lemma max_cases (a b : α) : max a b = a ∧ b ≤ a ∨ max a b = b ∧ a < b := @min_cases αᵒᵈ _ a b lemma min_eq_iff : min a b = c ↔ a = c ∧ a ≤ b ∨ b = c ∧ b ≤ a := begin @@ -81,8 +80,7 @@ begin simp [h] } end -lemma max_eq_iff : max a b = c ↔ a = c ∧ b ≤ a ∨ b = c ∧ a ≤ b := -@min_eq_iff (order_dual α) _ a b c +lemma max_eq_iff : max a b = c ↔ a = c ∧ b ≤ a ∨ b = c ∧ a ≤ b := @min_eq_iff αᵒᵈ _ a b c lemma min_lt_min_left_iff : min a c < min b c ↔ a < b ∧ a < c := by { simp_rw [lt_min_iff, min_lt_iff, or_iff_left (lt_irrefl _)], @@ -91,10 +89,8 @@ by { simp_rw [lt_min_iff, min_lt_iff, or_iff_left (lt_irrefl _)], lemma min_lt_min_right_iff : min a b < min a c ↔ b < c ∧ b < a := by simp_rw [min_comm a, min_lt_min_left_iff] -lemma max_lt_max_left_iff : max a c < max b c ↔ a < b ∧ c < b := -@min_lt_min_left_iff (order_dual α) _ _ _ _ -lemma max_lt_max_right_iff : max a b < max a c ↔ b < c ∧ a < c := -@min_lt_min_right_iff (order_dual α) _ _ _ _ +lemma max_lt_max_left_iff : max a c < max b c ↔ a < b ∧ c < b := @min_lt_min_left_iff αᵒᵈ _ _ _ _ +lemma max_lt_max_right_iff : max a b < max a c ↔ b < c ∧ a < c := @min_lt_min_right_iff αᵒᵈ _ _ _ _ /-- An instance asserting that `max a a = a` -/ instance max_idem : is_idempotent α max := by apply_instance -- short-circuit type class inference @@ -107,8 +103,7 @@ lemma min_lt_max : min a b < max a b ↔ a ≠ b := inf_lt_sup lemma max_lt_max (h₁ : a < c) (h₂ : b < d) : max a b < max c d := by simp [lt_max_iff, max_lt_iff, *] -lemma min_lt_min (h₁ : a < c) (h₂ : b < d) : min a b < min c d := -@max_lt_max (order_dual α) _ _ _ _ _ h₁ h₂ +lemma min_lt_min (h₁ : a < c) (h₂ : b < d) : min a b < min c d := @max_lt_max αᵒᵈ _ _ _ _ _ h₁ h₂ theorem min_right_comm (a b c : α) : min (min a b) c = min (min a c) b := right_comm min min_comm min_assoc a b c @@ -152,7 +147,7 @@ lemma min_rec {p : α → Prop} {x y : α} (hx : x ≤ y → p x) (hy : y ≤ x (λ h, (min_eq_right h).symm.subst (hy h)) lemma max_rec {p : α → Prop} {x y : α} (hx : y ≤ x → p x) (hy : x ≤ y → p y) : p (max x y) := -@min_rec (order_dual α) _ _ _ _ hx hy +@min_rec αᵒᵈ _ _ _ _ hx hy lemma min_rec' (p : α → Prop) {x y : α} (hx : p x) (hy : p y) : p (min x y) := min_rec (λ _, hx) (λ _, hy) @@ -164,7 +159,7 @@ theorem min_choice (a b : α) : min a b = a ∨ min a b = b := by cases le_total a b; simp * theorem max_choice (a b : α) : max a b = a ∨ max a b = b := -@min_choice (order_dual α) _ a b +@min_choice αᵒᵈ _ a b lemma le_of_max_le_left {a b c : α} (h : max a b ≤ c) : a ≤ c := le_trans (le_max_left _ _) h diff --git a/src/order/modular_lattice.lean b/src/order/modular_lattice.lean index 6901c7bc46e33..ce193145da41c 100644 --- a/src/order/modular_lattice.lean +++ b/src/order/modular_lattice.lean @@ -52,7 +52,7 @@ lemma inf_sup_assoc_of_le {x : α} (y : α) {z : α} (h : z ≤ x) : (x ⊓ y) ⊔ z = x ⊓ (y ⊔ z) := by rw [inf_comm, sup_comm, ← sup_inf_assoc_of_le y h, inf_comm, sup_comm] -instance : is_modular_lattice (order_dual α) := +instance : is_modular_lattice αᵒᵈ := ⟨λ x y z xz, le_of_eq (by { rw [inf_comm, sup_comm, eq_comm, inf_comm, sup_comm], convert sup_inf_assoc_of_le (order_dual.of_dual y) (order_dual.dual_le.2 xz) })⟩ @@ -60,7 +60,7 @@ variables {x y z : α} theorem is_modular_lattice.sup_inf_sup_assoc : (x ⊔ z) ⊓ (y ⊔ z) = ((x ⊔ z) ⊓ y) ⊔ z := -@is_modular_lattice.inf_sup_inf_assoc (order_dual α) _ _ _ _ _ +@is_modular_lattice.inf_sup_inf_assoc αᵒᵈ _ _ _ _ _ theorem eq_of_le_of_inf_le_of_sup_le (hxy : x ≤ y) (hinf : y ⊓ z ≤ x ⊓ z) (hsup : y ⊔ z ≤ x ⊔ z) : x = y := @@ -81,7 +81,7 @@ lt_of_le_of_ne (le_of_eq hsup.symm)) theorem inf_lt_inf_of_lt_of_sup_le_sup (hxy : x < y) (hinf : y ⊔ z ≤ x ⊔ z) : x ⊓ z < y ⊓ z := -@sup_lt_sup_of_lt_of_inf_le_inf (order_dual α) _ _ _ _ _ hxy hinf +@sup_lt_sup_of_lt_of_inf_le_inf αᵒᵈ _ _ _ _ _ hxy hinf /-- A generalization of the theorem that if `N` is a submodule of `M` and `N` and `M / N` are both Artinian, then `M` is Artinian. -/ @@ -119,9 +119,7 @@ theorem well_founded_gt_exact_sequence (hf : ∀ a, f₁ (f₂ a) = a ⊓ K) (hg : ∀ a, g₁ (g₂ a) = a ⊔ K) : well_founded ((>) : α → α → Prop) := -@well_founded_lt_exact_sequence - (order_dual α) _ _ (order_dual γ) (order_dual β) _ _ - h₂ h₁ K g₁ g₂ f₁ f₂ gi.dual gci.dual hg hf +@well_founded_lt_exact_sequence αᵒᵈ _ _ γᵒᵈ βᵒᵈ _ _ h₂ h₁ K g₁ g₂ f₁ f₂ gi.dual gci.dual hg hf /-- The diamond isomorphism between the intervals `[a ⊓ b, a]` and `[b, a ⊔ b]` -/ def inf_Icc_order_iso_Icc_sup (a b : α) : set.Icc (a ⊓ b) a ≃o set.Icc b (a ⊔ b) := diff --git a/src/order/monotone.lean b/src/order/monotone.lean index 8df550fd1916a..57f24e1e1f547 100644 --- a/src/order/monotone.lean +++ b/src/order/monotone.lean @@ -637,17 +637,14 @@ lemma monotone_nat_of_le_succ {f : ℕ → α} (hf : ∀ n, f n ≤ f (n + 1)) : monotone f := nat.rel_of_forall_rel_succ_of_le (≤) hf -lemma antitone_nat_of_succ_le {f : ℕ → α} (hf : ∀ n, f (n + 1) ≤ f n) : - antitone f := -@monotone_nat_of_le_succ (order_dual α) _ _ hf +lemma antitone_nat_of_succ_le {f : ℕ → α} (hf : ∀ n, f (n + 1) ≤ f n) : antitone f := +@monotone_nat_of_le_succ αᵒᵈ _ _ hf -lemma strict_mono_nat_of_lt_succ {f : ℕ → α} (hf : ∀ n, f n < f (n + 1)) : - strict_mono f := +lemma strict_mono_nat_of_lt_succ {f : ℕ → α} (hf : ∀ n, f n < f (n + 1)) : strict_mono f := nat.rel_of_forall_rel_succ_of_lt (<) hf -lemma strict_anti_nat_of_succ_lt {f : ℕ → α} (hf : ∀ n, f (n + 1) < f n) : - strict_anti f := -@strict_mono_nat_of_lt_succ (order_dual α) _ f hf +lemma strict_anti_nat_of_succ_lt {f : ℕ → α} (hf : ∀ n, f (n + 1) < f n) : strict_anti f := +@strict_mono_nat_of_lt_succ αᵒᵈ _ f hf lemma int.rel_of_forall_rel_succ_of_lt (r : β → β → Prop) [is_trans β r] {f : ℤ → β} (h : ∀ n, r (f n) (f (n + 1))) ⦃a b : ℤ⦄ (hab : a < b) : r (f a) (f b) := diff --git a/src/order/order_dual.lean b/src/order/order_dual.lean index a0a833081b658..b11185399d059 100644 --- a/src/order/order_dual.lean +++ b/src/order/order_dual.lean @@ -25,61 +25,38 @@ variables {α : Type u} {β : Type v} {γ : Type w} {r : α → α → Prop} namespace order_dual -instance [nontrivial α] : nontrivial (order_dual α) := by delta order_dual; assumption +instance [nontrivial α] : nontrivial αᵒᵈ := by delta order_dual; assumption /-- `to_dual` is the identity function to the `order_dual` of a linear order. -/ -def to_dual : α ≃ order_dual α := ⟨id, id, λ h, rfl, λ h, rfl⟩ +def to_dual : α ≃ αᵒᵈ := ⟨id, id, λ h, rfl, λ h, rfl⟩ /-- `of_dual` is the identity function from the `order_dual` of a linear order. -/ -def of_dual : order_dual α ≃ α := to_dual.symm +def of_dual : αᵒᵈ ≃ α := to_dual.symm @[simp] lemma to_dual_symm_eq : (@to_dual α).symm = of_dual := rfl - @[simp] lemma of_dual_symm_eq : (@of_dual α).symm = to_dual := rfl - -@[simp] lemma to_dual_of_dual (a : order_dual α) : to_dual (of_dual a) = a := rfl +@[simp] lemma to_dual_of_dual (a : αᵒᵈ) : to_dual (of_dual a) = a := rfl @[simp] lemma of_dual_to_dual (a : α) : of_dual (to_dual a) = a := rfl -@[simp] lemma to_dual_inj {a b : α} : - to_dual a = to_dual b ↔ a = b := iff.rfl - -@[simp] lemma to_dual_le_to_dual [has_le α] {a b : α} : - to_dual a ≤ to_dual b ↔ b ≤ a := iff.rfl - -@[simp] lemma to_dual_lt_to_dual [has_lt α] {a b : α} : - to_dual a < to_dual b ↔ b < a := iff.rfl - -@[simp] lemma of_dual_inj {a b : order_dual α} : - of_dual a = of_dual b ↔ a = b := iff.rfl +@[simp] lemma to_dual_inj {a b : α} : to_dual a = to_dual b ↔ a = b := iff.rfl +@[simp] lemma of_dual_inj {a b : αᵒᵈ} : of_dual a = of_dual b ↔ a = b := iff.rfl -@[simp] lemma of_dual_le_of_dual [has_le α] {a b : order_dual α} : - of_dual a ≤ of_dual b ↔ b ≤ a := iff.rfl +@[simp] lemma to_dual_le_to_dual [has_le α] {a b : α} : to_dual a ≤ to_dual b ↔ b ≤ a := iff.rfl +@[simp] lemma to_dual_lt_to_dual [has_lt α] {a b : α} : to_dual a < to_dual b ↔ b < a := iff.rfl +@[simp] lemma of_dual_le_of_dual [has_le α] {a b : αᵒᵈ} : of_dual a ≤ of_dual b ↔ b ≤ a := iff.rfl +@[simp] lemma of_dual_lt_of_dual [has_lt α] {a b : αᵒᵈ} : of_dual a < of_dual b ↔ b < a := iff.rfl -@[simp] lemma of_dual_lt_of_dual [has_lt α] {a b : order_dual α} : - of_dual a < of_dual b ↔ b < a := iff.rfl +lemma le_to_dual [has_le α] {a : αᵒᵈ} {b : α} : a ≤ to_dual b ↔ b ≤ of_dual a := iff.rfl +lemma lt_to_dual [has_lt α] {a : αᵒᵈ} {b : α} : a < to_dual b ↔ b < of_dual a := iff.rfl +lemma to_dual_le [has_le α] {a : α} {b : αᵒᵈ} : to_dual a ≤ b ↔ of_dual b ≤ a := iff.rfl +lemma to_dual_lt [has_lt α] {a : α} {b : αᵒᵈ} : to_dual a < b ↔ of_dual b < a := iff.rfl -lemma le_to_dual [has_le α] {a : order_dual α} {b : α} : - a ≤ to_dual b ↔ b ≤ of_dual a := iff.rfl - -lemma lt_to_dual [has_lt α] {a : order_dual α} {b : α} : - a < to_dual b ↔ b < of_dual a := iff.rfl - -lemma to_dual_le [has_le α] {a : α} {b : order_dual α} : - to_dual a ≤ b ↔ of_dual b ≤ a := iff.rfl - -lemma to_dual_lt [has_lt α] {a : α} {b : order_dual α} : - to_dual a < b ↔ of_dual b < a := iff.rfl - -/-- Recursor for `order_dual α`. -/ +/-- Recursor for `αᵒᵈ`. -/ @[elab_as_eliminator] -protected def rec {C : order_dual α → Sort*} (h₂ : Π (a : α), C (to_dual a)) : - Π (a : order_dual α), C a := h₂ - -@[simp] protected lemma «forall» {p : order_dual α → Prop} : (∀ a, p a) ↔ ∀ a, p (to_dual a) := -iff.rfl +protected def rec {C : αᵒᵈ → Sort*} (h₂ : Π (a : α), C (to_dual a)) : Π (a : αᵒᵈ), C a := h₂ -@[simp] protected lemma «exists» {p : order_dual α → Prop} : (∃ a, p a) ↔ ∃ a, p (to_dual a) := -iff.rfl +@[simp] protected lemma «forall» {p : αᵒᵈ → Prop} : (∀ a, p a) ↔ ∀ a, p (to_dual a) := iff.rfl +@[simp] protected lemma «exists» {p : αᵒᵈ → Prop} : (∃ a, p a) ↔ ∃ a, p (to_dual a) := iff.rfl end order_dual diff --git a/src/order/pfilter.lean b/src/order/pfilter.lean index 43b49eb367283..e421064d01b2a 100644 --- a/src/order/pfilter.lean +++ b/src/order/pfilter.lean @@ -16,7 +16,7 @@ a join-semilattice structure. - `order.pfilter P`: The type of nonempty, downward directed, upward closed subsets of `P`. This is dual to `order.ideal`, so it - simply wraps `order.ideal (order_dual P)`. + simply wraps `order.ideal Pᵒᵈ`. - `order.is_pfilter P`: a predicate for when a `set P` is a filter. @@ -43,11 +43,11 @@ variables {P : Type*} - downward directed - upward closed. -/ structure pfilter (P) [preorder P] := -(dual : ideal (order_dual P)) +(dual : ideal Pᵒᵈ) /-- A predicate for when a subset of `P` is a filter. -/ def is_pfilter [preorder P] (F : set P) : Prop := -@is_ideal (order_dual P) _ F +@is_ideal Pᵒᵈ _ F lemma is_pfilter.of_def [preorder P] {F : set P} (nonempty : F.nonempty) (directed : directed_on (≥) F) (mem_of_le : ∀ {x y : P}, x ≤ y → x ∈ F → y ∈ F) : is_pfilter F := diff --git a/src/order/rel_classes.lean b/src/order/rel_classes.lean index 8db3f3846d5bf..d2f5c9a52f45c 100644 --- a/src/order/rel_classes.lean +++ b/src/order/rel_classes.lean @@ -458,10 +458,10 @@ lemma transitive_lt [preorder α] : transitive (@has_lt.lt α _) := transitive_o lemma transitive_ge [preorder α] : transitive (@ge α _) := transitive_of_trans _ lemma transitive_gt [preorder α] : transitive (@gt α _) := transitive_of_trans _ -instance order_dual.is_total_le [has_le α] [is_total α (≤)] : is_total (order_dual α) (≤) := +instance order_dual.is_total_le [has_le α] [is_total α (≤)] : is_total αᵒᵈ (≤) := @is_total.swap α _ _ instance nat.lt.is_well_order : is_well_order ℕ (<) := ⟨nat.lt_wf⟩ -instance [linear_order α] [h : is_well_order α (<)] : is_well_order (order_dual α) (>) := h -instance [linear_order α] [h : is_well_order α (>)] : is_well_order (order_dual α) (<) := h +instance [linear_order α] [h : is_well_order α (<)] : is_well_order αᵒᵈ (>) := h +instance [linear_order α] [h : is_well_order α (>)] : is_well_order αᵒᵈ (<) := h diff --git a/src/order/rel_iso.lean b/src/order/rel_iso.lean index f92c19c5168cd..bbb073b837f78 100644 --- a/src/order/rel_iso.lean +++ b/src/order/rel_iso.lean @@ -74,7 +74,7 @@ lemma map_inf [semilattice_inf α] [linear_order β] lemma map_sup [semilattice_sup α] [linear_order β] [rel_hom_class F ((>) : β → β → Prop) ((>) : α → α → Prop)] (a : F) (m n : β) : a (m ⊔ n) = a m ⊔ a n := -@map_inf (order_dual α) (order_dual β) _ _ _ _ _ _ _ +@map_inf αᵒᵈ βᵒᵈ _ _ _ _ _ _ _ protected theorem is_irrefl [rel_hom_class F r s] (f : F) : ∀ [is_irrefl β s], is_irrefl α r | ⟨H⟩ := ⟨λ a h, H _ (map_rel f h)⟩ diff --git a/src/order/succ_pred/basic.lean b/src/order/succ_pred/basic.lean index 49640974bca81..a657e33acc0b5 100644 --- a/src/order/succ_pred/basic.lean +++ b/src/order/succ_pred/basic.lean @@ -72,14 +72,14 @@ variables {α : Type*} (le_pred_of_lt {a b} : a < b → a ≤ pred b) (le_of_pred_lt {a b} : pred a < b → a ≤ b) -instance [preorder α] [succ_order α] : pred_order (order_dual α) := +instance [preorder α] [succ_order α] : pred_order αᵒᵈ := { pred := to_dual ∘ succ_order.succ ∘ of_dual, pred_le := succ_order.le_succ, min_of_le_pred := λ _, succ_order.max_of_succ_le, le_pred_of_lt := λ a b h, succ_order.succ_le_of_lt h, le_of_pred_lt := λ a b, succ_order.le_of_lt_succ } -instance [preorder α] [pred_order α] : succ_order (order_dual α) := +instance [preorder α] [pred_order α] : succ_order αᵒᵈ := { succ := to_dual ∘ pred_order.pred ∘ of_dual, le_succ := pred_order.pred_le, max_of_succ_le := λ _, pred_order.min_of_le_pred, @@ -532,8 +532,8 @@ variables [order_bot α] @[simp] lemma pred_bot : pred (⊥ : α) = ⊥ := is_min_bot.pred_eq -@[simp] lemma le_pred_iff_eq_bot : a ≤ pred a ↔ a = ⊥ := @succ_le_iff_eq_top (order_dual α) _ _ _ _ -@[simp] lemma pred_lt_iff_ne_bot : pred a < a ↔ a ≠ ⊥ := @lt_succ_iff_ne_top (order_dual α) _ _ _ _ +@[simp] lemma le_pred_iff_eq_bot : a ≤ pred a ↔ a = ⊥ := @succ_le_iff_eq_top αᵒᵈ _ _ _ _ +@[simp] lemma pred_lt_iff_ne_bot : pred a < a ↔ a ≠ ⊥ := @lt_succ_iff_ne_top αᵒᵈ _ _ _ _ end order_bot @@ -880,8 +880,7 @@ variables [preorder α] section succ_order variables [succ_order α] [is_succ_archimedean α] {a b : α} -instance : is_pred_archimedean (order_dual α) := -⟨λ a b h, by convert exists_succ_iterate_of_le h.of_dual⟩ +instance : is_pred_archimedean αᵒᵈ := ⟨λ a b h, by convert exists_succ_iterate_of_le h.of_dual⟩ lemma has_le.le.exists_succ_iterate (h : a ≤ b) : ∃ n, succ^[n] a = b := exists_succ_iterate_of_le h @@ -915,23 +914,22 @@ end succ_order section pred_order variables [pred_order α] [is_pred_archimedean α] {a b : α} -instance : is_succ_archimedean (order_dual α) := -⟨λ a b h, by convert exists_pred_iterate_of_le h.of_dual⟩ +instance : is_succ_archimedean αᵒᵈ := ⟨λ a b h, by convert exists_pred_iterate_of_le h.of_dual⟩ lemma has_le.le.exists_pred_iterate (h : a ≤ b) : ∃ n, pred^[n] b = a := exists_pred_iterate_of_le h lemma exists_pred_iterate_iff_le : (∃ n, pred^[n] b = a) ↔ a ≤ b := -@exists_succ_iterate_iff_le (order_dual α) _ _ _ _ _ +@exists_succ_iterate_iff_le αᵒᵈ _ _ _ _ _ /-- Induction principle on a type with a `pred_order` for all elements below a given element `m`. -/ @[elab_as_eliminator] lemma pred.rec {P : α → Prop} {m : α} (h0 : P m) (h1 : ∀ n, n ≤ m → P n → P (pred n)) ⦃n : α⦄ (hmn : n ≤ m) : P n := -@succ.rec (order_dual α) _ _ _ _ _ h0 h1 _ hmn +@succ.rec αᵒᵈ _ _ _ _ _ h0 h1 _ hmn lemma pred.rec_iff {p : α → Prop} (hsucc : ∀ a, p a ↔ p (pred a)) {a b : α} (h : a ≤ b) : p a ↔ p b := -(@succ.rec_iff (order_dual α) _ _ _ _ hsucc _ _ h).symm +(@succ.rec_iff αᵒᵈ _ _ _ _ hsucc _ _ h).symm end pred_order end preorder @@ -983,7 +981,7 @@ end⟩ @[priority 100] instance is_well_order.to_is_succ_archimedean [h : is_well_order α (>)] [succ_order α] : is_succ_archimedean α := -by convert @order_dual.is_succ_archimedean (order_dual α) _ _ _ +by convert @order_dual.is_succ_archimedean αᵒᵈ _ _ _ end is_well_order diff --git a/src/order/succ_pred/relation.lean b/src/order/succ_pred/relation.lean index 0767f09851d15..15a145eeb1af0 100644 --- a/src/order/succ_pred/relation.lean +++ b/src/order/succ_pred/relation.lean @@ -86,25 +86,25 @@ variables {α : Type*} [partial_order α] [pred_order α] [is_pred_archimedean for all `i` between `n` and `m`. -/ lemma refl_trans_gen_of_pred_of_ge (r : α → α → Prop) {n m : α} (h : ∀ i ∈ Ioc m n, r i (pred i)) (hnm : m ≤ n) : refl_trans_gen r n m := -@refl_trans_gen_of_succ_of_le (order_dual α) _ _ _ r n m (λ x hx, h x ⟨hx.2, hx.1⟩) hnm +@refl_trans_gen_of_succ_of_le αᵒᵈ _ _ _ r n m (λ x hx, h x ⟨hx.2, hx.1⟩) hnm /-- For `n ≤ m`, `(n, m)` is in the reflexive-transitive closure of `~` if `pred i ~ i` for all `i` between `n` and `m`. -/ lemma refl_trans_gen_of_pred_of_le (r : α → α → Prop) {n m : α} (h : ∀ i ∈ Ioc n m, r (pred i) i) (hmn : n ≤ m) : refl_trans_gen r n m := -@refl_trans_gen_of_succ_of_ge (order_dual α) _ _ _ r n m (λ x hx, h x ⟨hx.2, hx.1⟩) hmn +@refl_trans_gen_of_succ_of_ge αᵒᵈ _ _ _ r n m (λ x hx, h x ⟨hx.2, hx.1⟩) hmn /-- For `m < n`, `(n, m)` is in the transitive closure of a relation `~` for `n ≠ m` if `i ~ pred i` for all `i` between `n` and `m`. -/ lemma trans_gen_of_pred_of_gt (r : α → α → Prop) {n m : α} (h : ∀ i ∈ Ioc m n, r i (pred i)) (hnm : m < n) : trans_gen r n m := -@trans_gen_of_succ_of_lt (order_dual α) _ _ _ r _ _ (λ x hx, h x ⟨hx.2, hx.1⟩) hnm +@trans_gen_of_succ_of_lt αᵒᵈ _ _ _ r _ _ (λ x hx, h x ⟨hx.2, hx.1⟩) hnm /-- For `n < m`, `(n, m)` is in the transitive closure of a relation `~` for `n ≠ m` if `pred i ~ i` for all `i` between `n` and `m`. -/ lemma trans_gen_of_pred_of_lt (r : α → α → Prop) {n m : α} (h : ∀ i ∈ Ioc n m, r (pred i) i) (hmn : n < m) : trans_gen r n m := -@trans_gen_of_succ_of_gt (order_dual α) _ _ _ r _ _ (λ x hx, h x ⟨hx.2, hx.1⟩) hmn +@trans_gen_of_succ_of_gt αᵒᵈ _ _ _ r _ _ (λ x hx, h x ⟨hx.2, hx.1⟩) hmn end partial_pred @@ -115,7 +115,7 @@ variables {α : Type*} [linear_order α] [pred_order α] [is_pred_archimedean α for all `i` between `n` and `m`. -/ lemma refl_trans_gen_of_pred (r : α → α → Prop) {n m : α} (h1 : ∀ i ∈ Ioc m n, r i (pred i)) (h2 : ∀ i ∈ Ioc n m, r (pred i) i) : refl_trans_gen r n m := -@refl_trans_gen_of_succ (order_dual α) _ _ _ r n m (λ x hx, h1 x ⟨hx.2, hx.1⟩) +@refl_trans_gen_of_succ αᵒᵈ _ _ _ r n m (λ x hx, h1 x ⟨hx.2, hx.1⟩) (λ x hx, h2 x ⟨hx.2, hx.1⟩) /-- For `n ≠ m`, `(n, m)` is in the transitive closure of a relation `~` if `i ~ pred i` and @@ -123,14 +123,14 @@ lemma refl_trans_gen_of_pred (r : α → α → Prop) {n m : α} lemma trans_gen_of_pred_of_ne (r : α → α → Prop) {n m : α} (h1 : ∀ i ∈ Ioc m n, r i (pred i)) (h2 : ∀ i ∈ Ioc n m, r (pred i) i) (hnm : n ≠ m) : trans_gen r n m := -@trans_gen_of_succ_of_ne (order_dual α) _ _ _ r n m (λ x hx, h1 x ⟨hx.2, hx.1⟩) +@trans_gen_of_succ_of_ne αᵒᵈ _ _ _ r n m (λ x hx, h1 x ⟨hx.2, hx.1⟩) (λ x hx, h2 x ⟨hx.2, hx.1⟩) hnm /-- `(n, m)` is in the transitive closure of a reflexive relation `~` if `i ~ pred i` and `pred i ~ i` for all `i` between `n` and `m`. -/ lemma trans_gen_of_pred_of_reflexive (r : α → α → Prop) {n m : α} (hr : reflexive r) (h1 : ∀ i ∈ Ioc m n, r i (pred i)) (h2 : ∀ i ∈ Ioc n m, r (pred i) i) : trans_gen r n m := -@trans_gen_of_succ_of_reflexive (order_dual α) _ _ _ r n m hr (λ x hx, h1 x ⟨hx.2, hx.1⟩) +@trans_gen_of_succ_of_reflexive αᵒᵈ _ _ _ r n m hr (λ x hx, h1 x ⟨hx.2, hx.1⟩) (λ x hx, h2 x ⟨hx.2, hx.1⟩) end linear_pred diff --git a/src/order/upper_lower.lean b/src/order/upper_lower.lean index 18ca128ad0824..9ece02f8ee12f 100644 --- a/src/order/upper_lower.lean +++ b/src/order/upper_lower.lean @@ -109,9 +109,9 @@ lemma is_lower_set_sInter {S : set (set α)} (hf : ∀ s ∈ S, is_lower_set s) iff.rfl @[simp] lemma is_upper_set_preimage_of_dual_iff : is_upper_set (of_dual ⁻¹' s) ↔ is_lower_set s := iff.rfl -@[simp] lemma is_lower_set_preimage_to_dual_iff {s : set (order_dual α)} : +@[simp] lemma is_lower_set_preimage_to_dual_iff {s : set αᵒᵈ} : is_lower_set (to_dual ⁻¹' s) ↔ is_upper_set s := iff.rfl -@[simp] lemma is_upper_set_preimage_to_dual_iff {s : set (order_dual α)} : +@[simp] lemma is_upper_set_preimage_to_dual_iff {s : set αᵒᵈ} : is_upper_set (to_dual ⁻¹' s) ↔ is_lower_set s := iff.rfl alias is_lower_set_preimage_of_dual_iff ↔ _ is_upper_set.of_dual @@ -383,7 +383,7 @@ variables [semilattice_sup α] @[simp] lemma Ici_sup (a b : α) : Ici (a ⊔ b) = Ici a ⊓ Ici b := ext Ici_inter_Ici.symm /-- `upper_set.Ici` as a `sup_hom`. -/ -def Ici_sup_hom : sup_hom α (order_dual $ upper_set α) := ⟨Ici, Ici_sup⟩ +def Ici_sup_hom : sup_hom α (upper_set α)ᵒᵈ := ⟨Ici, Ici_sup⟩ @[simp] lemma Ici_sup_hom_apply (a : α) : Ici_sup_hom a = to_dual (Ici a) := rfl @@ -402,7 +402,7 @@ set_like.ext $ λ c, by simp only [mem_Ici_iff, mem_infi_iff, supr_le_iff] by simp_rw Ici_supr /-- `upper_set.Ici` as a `Sup_hom`. -/ -def Ici_Sup_hom : Sup_hom α (order_dual $ upper_set α) := +def Ici_Sup_hom : Sup_hom α (upper_set α)ᵒᵈ := ⟨Ici, λ s, (Ici_Sup s).trans Inf_image.symm⟩ @[simp] lemma Ici_Sup_hom_apply (a : α) : Ici_Sup_hom a = to_dual (Ici a) := rfl diff --git a/src/order/well_founded.lean b/src/order/well_founded.lean index 2642c4ba58431..2323747cf163d 100644 --- a/src/order/well_founded.lean +++ b/src/order/well_founded.lean @@ -94,7 +94,7 @@ by simp only [eq_iff_not_lt_of_le, well_founded_iff_has_min] theorem well_founded_iff_has_min' [partial_order α] : (well_founded (has_lt.lt : α → α → Prop)) ↔ ∀ (p : set α), p.nonempty → ∃ m ∈ p, ∀ x ∈ p, x ≤ m → x = m := -@well_founded_iff_has_max' (order_dual α) _ +@well_founded_iff_has_max' αᵒᵈ _ open set /-- The supremum of a bounded, well-founded order -/ diff --git a/src/order/well_founded_set.lean b/src/order/well_founded_set.lean index c3ca197b9ac11..d425dd2f43372 100644 --- a/src/order/well_founded_set.lean +++ b/src/order/well_founded_set.lean @@ -131,8 +131,7 @@ end has_lt section partial_order variables [partial_order α] {s t : set α} {a : α} -theorem is_wf_iff_no_descending_seq : - is_wf s ↔ ∀ (f : (order_dual ℕ) ↪o α), ¬ (range f) ⊆ s := +theorem is_wf_iff_no_descending_seq : is_wf s ↔ ∀ f : ℕᵒᵈ ↪o α, ¬ (range f) ⊆ s := begin haveI : is_strict_order α (λ (a b : α), a < b ∧ a ∈ s ∧ b ∈ s) := { to_is_irrefl := ⟨λ x con, lt_irrefl x con.1⟩, diff --git a/src/order/zorn.lean b/src/order/zorn.lean index ac7aa1ffe64b5..2804ab41623b0 100644 --- a/src/order/zorn.lean +++ b/src/order/zorn.lean @@ -168,13 +168,12 @@ zorn_nonempty_partial_order₀ _ (λ c cS hc y yc, H _ cS hc ⟨y, yc⟩) _ hx lemma zorn_superset (S : set (set α)) (h : ∀ c ⊆ S, is_chain (⊆) c → ∃ lb ∈ S, ∀ s ∈ c, lb ⊆ s) : ∃ m ∈ S, ∀ a ∈ S, a ⊆ m → a = m := -@zorn_partial_order₀ (order_dual (set α)) _ S $ λ c cS hc, h c cS hc.symm +@zorn_partial_order₀ (set α)ᵒᵈ _ S $ λ c cS hc, h c cS hc.symm lemma zorn_superset_nonempty (S : set (set α)) (H : ∀ c ⊆ S, is_chain (⊆) c → c.nonempty → ∃ lb ∈ S, ∀ s ∈ c, lb ⊆ s) (x) (hx : x ∈ S) : ∃ m ∈ S, m ⊆ x ∧ ∀ a ∈ S, a ⊆ m → a = m := -@zorn_nonempty_partial_order₀ (order_dual (set α)) _ S (λ c cS hc y yc, H _ cS - hc.symm ⟨y, yc⟩) _ hx +@zorn_nonempty_partial_order₀ (set α)ᵒᵈ _ S (λ c cS hc y yc, H _ cS hc.symm ⟨y, yc⟩) _ hx /-- Every chain is contained in a maximal chain. This generalizes Hausdorff's maximality principle. -/ diff --git a/src/ring_theory/artinian.lean b/src/ring_theory/artinian.lean index 3add0c5b5d499..4f32cb64d8db4 100644 --- a/src/ring_theory/artinian.lean +++ b/src/ring_theory/artinian.lean @@ -201,12 +201,12 @@ set_has_minimal_iff_artinian.mpr ‹_› a ha /-- A module is Artinian iff every decreasing chain of submodules stabilizes. -/ theorem monotone_stabilizes_iff_artinian : - (∀ (f : ℕ →o order_dual (submodule R M)), ∃ n, ∀ m, n ≤ m → f n = f m) + (∀ (f : ℕ →o (submodule R M)ᵒᵈ), ∃ n, ∀ m, n ≤ m → f n = f m) ↔ is_artinian R M := by rw [is_artinian_iff_well_founded]; - exact (well_founded.monotone_chain_condition (order_dual (submodule R M))).symm + exact (well_founded.monotone_chain_condition (submodule R M)ᵒᵈ).symm -theorem is_artinian.monotone_stabilizes [is_artinian R M] (f : ℕ →o order_dual (submodule R M)) : +theorem is_artinian.monotone_stabilizes [is_artinian R M] (f : ℕ →o (submodule R M)ᵒᵈ) : ∃ n, ∀ m, n ≤ m → f n = f m := monotone_stabilizes_iff_artinian.mpr ‹_› f @@ -399,7 +399,7 @@ variables {R : Type*} [comm_ring R] [is_artinian_ring R] lemma is_nilpotent_jacobson_bot : is_nilpotent (ideal.jacobson (⊥ : ideal R)) := begin let Jac := ideal.jacobson (⊥ : ideal R), - let f : ℕ →o order_dual (ideal R) := ⟨λ n, Jac ^ n, λ _ _ h, ideal.pow_le_pow h⟩, + let f : ℕ →o (ideal R)ᵒᵈ := ⟨λ n, Jac ^ n, λ _ _ h, ideal.pow_le_pow h⟩, obtain ⟨n, hn⟩ : ∃ n, ∀ m, n ≤ m → Jac ^ n = Jac ^ m := is_artinian.monotone_stabilizes f, refine ⟨n, _⟩, let J : ideal R := annihilator (Jac ^ n), diff --git a/src/ring_theory/nullstellensatz.lean b/src/ring_theory/nullstellensatz.lean index e182530ddc107..91682491d504a 100644 --- a/src/ring_theory/nullstellensatz.lean +++ b/src/ring_theory/nullstellensatz.lean @@ -76,7 +76,7 @@ lemma zero_locus_vanishing_ideal_le (V : set (σ → k)) : λ V hV p hp, hp V hV theorem zero_locus_vanishing_ideal_galois_connection : - @galois_connection (ideal (mv_polynomial σ k)) (order_dual (set (σ → k))) _ _ + @galois_connection (ideal (mv_polynomial σ k)) (set (σ → k))ᵒᵈ _ _ zero_locus vanishing_ideal := λ I V, ⟨λ h, le_trans (le_vanishing_ideal_zero_locus I) (vanishing_ideal_anti_mono h), λ h, le_trans (zero_locus_anti_mono h) (zero_locus_vanishing_ideal_le V)⟩ diff --git a/src/ring_theory/valuation/basic.lean b/src/ring_theory/valuation/basic.lean index 4375641751d0c..e966286baa593 100644 --- a/src/ring_theory/valuation/basic.lean +++ b/src/ring_theory/valuation/basic.lean @@ -47,7 +47,7 @@ on R / J = `ideal.quotient J` is `on_quot v h`. ## Implementation Details -`add_valuation R Γ₀` is implemented as `valuation R (multiplicative (order_dual Γ₀))`. +`add_valuation R Γ₀` is implemented as `valuation R (multiplicative Γ₀)ᵒᵈ`. ## TODO @@ -504,7 +504,7 @@ variables (R) [ring R] (Γ₀ : Type*) [linear_ordered_add_comm_monoid_with_top /-- The type of `Γ₀`-valued additive valuations on `R`. -/ @[nolint has_inhabited_instance] -def add_valuation := valuation R (multiplicative (order_dual Γ₀)) +def add_valuation := valuation R (multiplicative Γ₀ᵒᵈ) end add_monoid @@ -529,8 +529,7 @@ section variables (f : R → Γ₀) (h0 : f 0 = ⊤) (h1 : f 1 = 0) variables (hadd : ∀ x y, min (f x) (f y) ≤ f (x + y)) (hmul : ∀ x y, f (x * y) = f x + f y) -/-- An alternate constructor of `add_valuation`, that doesn't reference - `multiplicative (order_dual Γ₀)` -/ +/-- An alternate constructor of `add_valuation`, that doesn't reference `multiplicative Γ₀ᵒᵈ` -/ def of : add_valuation R Γ₀ := { to_fun := f, map_one' := h1, @@ -545,7 +544,7 @@ theorem of_apply : (of f h0 h1 hadd hmul) r = f r := rfl /-- The `valuation` associated to an `add_valuation` (useful if the latter is constructed using `add_valuation.of`). -/ -def valuation : valuation R (multiplicative (order_dual Γ₀)) := v +def valuation : valuation R (multiplicative Γ₀ᵒᵈ) := v @[simp] lemma valuation_apply (r : R) : v.valuation r = multiplicative.of_add (order_dual.to_dual (v r)) := rfl diff --git a/src/ring_theory/valuation/valuation_subring.lean b/src/ring_theory/valuation/valuation_subring.lean index 83bfc7285a718..f814c27e93924 100644 --- a/src/ring_theory/valuation/valuation_subring.lean +++ b/src/ring_theory/valuation/valuation_subring.lean @@ -284,8 +284,7 @@ def prime_spectrum_equiv : /-- An ordered variant of `prime_spectrum_equiv`. -/ @[simps] -def prime_spectrum_order_equiv : - order_dual (prime_spectrum A) ≃o { S | A ≤ S } := +def prime_spectrum_order_equiv : (prime_spectrum A)ᵒᵈ ≃o {S | A ≤ S} := { map_rel_iff' := λ P Q, ⟨ λ h, begin have := ideal_of_le_le_of_le A _ _ _ _ h, diff --git a/src/topology/algebra/order/basic.lean b/src/topology/algebra/order/basic.lean index f1cb85b76af35..5894f7a29336b 100644 --- a/src/topology/algebra/order/basic.lean +++ b/src/topology/algebra/order/basic.lean @@ -90,17 +90,15 @@ generally, and suffices to derive many interesting properties relating order and class order_closed_topology (α : Type*) [topological_space α] [preorder α] : Prop := (is_closed_le' : is_closed {p:α×α | p.1 ≤ p.2}) -instance : Π [topological_space α], topological_space (order_dual α) := id +instance : Π [topological_space α], topological_space αᵒᵈ := id -instance [topological_space α] [h : first_countable_topology α] : - first_countable_topology (order_dual α) := h +instance [topological_space α] [h : first_countable_topology α] : first_countable_topology αᵒᵈ := h -instance [topological_space α] [h : second_countable_topology α] : - second_countable_topology (order_dual α) := h +instance [topological_space α] [h : second_countable_topology α] : second_countable_topology αᵒᵈ := +h @[to_additive] -instance [topological_space α] [has_mul α] [h : has_continuous_mul α] : - has_continuous_mul (order_dual α) := h +instance [topological_space α] [has_mul α] [h : has_continuous_mul α] : has_continuous_mul αᵒᵈ := h lemma dense.order_dual [topological_space α] {s : set α} (hs : dense s) : dense (order_dual.of_dual ⁻¹' s) := hs @@ -140,7 +138,7 @@ is_closed_le continuous_const continuous_id lemma is_closed_Ici {a : α} : is_closed (Ici a) := is_closed_ge' a -instance : order_closed_topology (order_dual α) := +instance : order_closed_topology αᵒᵈ := ⟨(@order_closed_topology.is_closed_le' α _ _ _).preimage continuous_swap⟩ lemma is_closed_Icc {a b : α} : is_closed (Icc a b) := @@ -534,8 +532,7 @@ end lemma frontier_Iic_subset (a : α) : frontier (Iic a) ⊆ {a} := frontier_le_subset_eq (@continuous_id α _) continuous_const -lemma frontier_Ici_subset (a : α) : frontier (Ici a) ⊆ {a} := -@frontier_Iic_subset (order_dual α) _ _ _ _ +lemma frontier_Ici_subset (a : α) : frontier (Ici a) ⊆ {a} := @frontier_Iic_subset αᵒᵈ _ _ _ _ lemma frontier_lt_subset_eq (hf : continuous f) (hg : continuous g) : frontier {b | f b < g b} ⊆ {b | f b = g b} := @@ -581,7 +578,7 @@ by { simp only [min_def], exact hf.if_le hg hf hg (λ x, id) } @[continuity] lemma continuous.max (hf : continuous f) (hg : continuous g) : continuous (λb, max (f b) (g b)) := -@continuous.min (order_dual α) _ _ _ _ _ _ _ hf hg +@continuous.min αᵒᵈ _ _ _ _ _ _ _ hf hg end @@ -646,7 +643,7 @@ end /-- A compact set is bounded above -/ lemma is_compact.bdd_above {s : set α} (hs : is_compact s) : bdd_above s := -@is_compact.bdd_below (order_dual α) _ _ _ _ _ hs +@is_compact.bdd_below αᵒᵈ _ _ _ _ _ hs /-- A continuous function is bounded below on a compact set. -/ lemma is_compact.bdd_below_image {f : β → α} {K : set β} @@ -656,7 +653,7 @@ lemma is_compact.bdd_below_image {f : β → α} {K : set β} /-- A continuous function is bounded above on a compact set. -/ lemma is_compact.bdd_above_image {f : β → α} {K : set β} (hK : is_compact K) (hf : continuous_on f K) : bdd_above (f '' K) := -@is_compact.bdd_below_image (order_dual α) _ _ _ _ _ _ _ _ hK hf +@is_compact.bdd_below_image αᵒᵈ _ _ _ _ _ _ _ _ hK hf /-- A continuous function with compact support is bounded below. -/ @[to_additive /-" A continuous function with compact support is bounded below. "-/] @@ -669,7 +666,7 @@ lemma continuous.bdd_below_range_of_has_compact_mul_support [has_one α] {f : β lemma continuous.bdd_above_range_of_has_compact_mul_support [has_one α] {f : β → α} (hf : continuous f) (h : has_compact_mul_support f) : bdd_above (range f) := -@continuous.bdd_below_range_of_has_compact_mul_support (order_dual α) _ _ _ _ _ _ _ _ hf h +@continuous.bdd_below_range_of_has_compact_mul_support αᵒᵈ _ _ _ _ _ _ _ _ hf h end linear_order @@ -711,7 +708,7 @@ generate_from {s : set α | ∃ (a : α), s = {b : α | a < b} ∨ s = {b : α | section order_topology instance {α : Type*} [topological_space α] [partial_order α] [order_topology α] : - order_topology (order_dual α) := + order_topology αᵒᵈ := ⟨by convert @order_topology.topology_eq_generate_intervals α _ _ _; conv in (_ ∨ _) { rw or.comm }; refl⟩ @@ -944,7 +941,7 @@ lemma nhds_top_basis [topological_space α] [linear_order α] [order_top α] [or lemma nhds_bot_basis [topological_space α] [linear_order α] [order_bot α] [order_topology α] [nontrivial α] : (𝓝 ⊥).has_basis (λ a : α, ⊥ < a) (λ a : α, Iio a) := -@nhds_top_basis (order_dual α) _ _ _ _ _ +@nhds_top_basis αᵒᵈ _ _ _ _ _ lemma nhds_top_basis_Ici [topological_space α] [linear_order α] [order_top α] [order_topology α] [nontrivial α] [densely_ordered α] : @@ -956,7 +953,7 @@ nhds_top_basis.to_has_basis lemma nhds_bot_basis_Iic [topological_space α] [linear_order α] [order_bot α] [order_topology α] [nontrivial α] [densely_ordered α] : (𝓝 ⊥).has_basis (λ a : α, ⊥ < a) Iic := -@nhds_top_basis_Ici (order_dual α) _ _ _ _ _ _ +@nhds_top_basis_Ici αᵒᵈ _ _ _ _ _ _ lemma tendsto_nhds_top_mono [topological_space β] [partial_order β] [order_top β] [order_topology β] {l : filter α} {f g : α → β} (hf : tendsto f l (𝓝 ⊤)) (hg : f ≤ᶠ[l] g) : @@ -970,7 +967,7 @@ end lemma tendsto_nhds_bot_mono [topological_space β] [partial_order β] [order_bot β] [order_topology β] {l : filter α} {f g : α → β} (hf : tendsto f l (𝓝 ⊥)) (hg : g ≤ᶠ[l] f) : tendsto g l (𝓝 ⊥) := -@tendsto_nhds_top_mono α (order_dual β) _ _ _ _ _ _ _ hf hg +@tendsto_nhds_top_mono α βᵒᵈ _ _ _ _ _ _ _ hf hg lemma tendsto_nhds_top_mono' [topological_space β] [partial_order β] [order_top β] [order_topology β] {l : filter α} {f g : α → β} (hf : tendsto f l (𝓝 ⊤)) (hg : f ≤ g) : @@ -1185,7 +1182,7 @@ lemma pi_Iio_mem_nhds' (ha : ∀ i, x' i < a' i) : Iio a' ∈ 𝓝 x' := pi_Iio_mem_nhds ha lemma pi_Ioi_mem_nhds (ha : ∀ i, a i < x i) : Ioi a ∈ 𝓝 x := -@pi_Iio_mem_nhds ι (λ i, order_dual (π i)) _ _ _ _ _ _ _ ha +@pi_Iio_mem_nhds ι (λ i, (π i)ᵒᵈ) _ _ _ _ _ _ _ ha lemma pi_Ioi_mem_nhds' (ha : ∀ i, a' i < x' i) : Ioi a' ∈ 𝓝 x' := pi_Ioi_mem_nhds ha @@ -1234,13 +1231,11 @@ end 𝓝 x ⊓ at_top = ⊥ := disjoint_iff.1 (disjoint_nhds_at_top x) -lemma disjoint_nhds_at_bot [no_min_order α] (x : α) : - disjoint (𝓝 x) at_bot := -@disjoint_nhds_at_top (order_dual α) _ _ _ _ x +lemma disjoint_nhds_at_bot [no_min_order α] (x : α) : disjoint (𝓝 x) at_bot := +@disjoint_nhds_at_top αᵒᵈ _ _ _ _ x -@[simp] lemma inf_nhds_at_bot [no_min_order α] (x : α) : - 𝓝 x ⊓ at_bot = ⊥ := -@inf_nhds_at_top (order_dual α) _ _ _ _ x +@[simp] lemma inf_nhds_at_bot [no_min_order α] (x : α) : 𝓝 x ⊓ at_bot = ⊥ := +@inf_nhds_at_top αᵒᵈ _ _ _ _ x lemma not_tendsto_nhds_of_tendsto_at_top [no_max_order α] {F : filter β} [ne_bot F] {f : β → α} (hf : tendsto f F at_top) (x : α) : @@ -1469,8 +1464,8 @@ with `l < a`. -/ lemma mem_nhds_within_Iic_iff_exists_Icc_subset' [no_min_order α] [densely_ordered α] {a : α} {s : set α} : s ∈ 𝓝[≤] a ↔ ∃l ∈ Iio a, Icc l a ⊆ s := begin - convert @mem_nhds_within_Ici_iff_exists_Icc_subset' (order_dual α) _ _ _ _ _ _ _, - simp_rw (show ∀ u : order_dual α, @Icc (order_dual α) _ a u = @Icc α _ u a, from λ u, dual_Icc), + convert @mem_nhds_within_Ici_iff_exists_Icc_subset' αᵒᵈ _ _ _ _ _ _ _, + simp_rw (show ∀ u : αᵒᵈ, @Icc αᵒᵈ _ a u = @Icc α _ u a, from λ u, dual_Icc), refl, end @@ -1661,7 +1656,7 @@ end and `g` tends to `at_bot` then `f + g` tends to `at_bot`. -/ lemma filter.tendsto.add_at_bot {C : α} (hf : tendsto f l (𝓝 C)) (hg : tendsto g l at_bot) : tendsto (λ x, f x + g x) l at_bot := -@filter.tendsto.add_at_top (order_dual α) _ _ _ _ _ _ _ _ hf hg +@filter.tendsto.add_at_top αᵒᵈ _ _ _ _ _ _ _ _ hf hg /-- In a linearly ordered additive commutative group with the order topology, if `f` tends to `at_top` and `g` tends to `C` then `f + g` tends to `at_top`. -/ @@ -1984,7 +1979,7 @@ lemma is_lub.frequently_nhds_mem {a : α} {s : set α} (ha : is_lub s a) (hs : s lemma is_glb.frequently_mem {a : α} {s : set α} (ha : is_glb s a) (hs : s.nonempty) : ∃ᶠ x in 𝓝[≥] a, x ∈ s := -@is_lub.frequently_mem (order_dual α) _ _ _ _ _ ha hs +@is_lub.frequently_mem αᵒᵈ _ _ _ _ _ ha hs lemma is_glb.frequently_nhds_mem {a : α} {s : set α} (ha : is_glb s a) (hs : s.nonempty) : ∃ᶠ x in 𝓝 a, x ∈ s := @@ -2004,7 +1999,7 @@ mem_closure_iff_nhds_within_ne_bot.1 (ha.mem_closure hs) lemma is_glb.nhds_within_ne_bot : ∀ {a : α} {s : set α}, is_glb s a → s.nonempty → ne_bot (𝓝[s] a) := -@is_lub.nhds_within_ne_bot (order_dual α) _ _ _ +@is_lub.nhds_within_ne_bot αᵒᵈ _ _ _ lemma is_lub_of_mem_nhds {s : set α} {a : α} {f : filter α} (hsa : a ∈ upper_bounds s) (hsf : s ∈ f) [ne_bot (f ⊓ 𝓝 a)] : is_lub s a := @@ -2026,11 +2021,11 @@ end lemma is_glb_of_mem_nhds : ∀ {s : set α} {a : α} {f : filter α}, a ∈ lower_bounds s → s ∈ f → ne_bot (f ⊓ 𝓝 a) → is_glb s a := -@is_lub_of_mem_nhds (order_dual α) _ _ _ +@is_lub_of_mem_nhds αᵒᵈ _ _ _ lemma is_glb_of_mem_closure {s : set α} {a : α} (hsa : a ∈ lower_bounds s) (hsf : a ∈ closure s) : is_glb s a := -@is_lub_of_mem_closure (order_dual α) _ _ _ s a hsa hsf +@is_lub_of_mem_closure αᵒᵈ _ _ _ s a hsa hsf lemma is_lub.mem_upper_bounds_of_tendsto [preorder γ] [topological_space γ] [order_closed_topology γ] {f : α → γ} {s : set α} {a : α} {b : γ} @@ -2060,7 +2055,7 @@ lemma is_glb.mem_lower_bounds_of_tendsto [preorder γ] [topological_space γ] [order_closed_topology γ] {f : α → γ} {s : set α} {a : α} {b : γ} (hf : monotone_on f s) (ha : is_glb s a) (hb : tendsto f (𝓝[s] a) (𝓝 b)) : b ∈ lower_bounds (f '' s) := -@is_lub.mem_upper_bounds_of_tendsto (order_dual α) (order_dual γ) _ _ _ _ _ _ _ _ _ _ hf.dual ha hb +@is_lub.mem_upper_bounds_of_tendsto αᵒᵈ γᵒᵈ _ _ _ _ _ _ _ _ _ _ hf.dual ha hb -- For a version of this theorem in which the convergence considered on the domain `α` is as -- `x : α` tends to negative infinity, rather than tending to a point `x` in `α`, see @@ -2069,31 +2064,31 @@ lemma is_glb.is_glb_of_tendsto [preorder γ] [topological_space γ] [order_closed_topology γ] {f : α → γ} {s : set α} {a : α} {b : γ} (hf : monotone_on f s) : is_glb s a → s.nonempty → tendsto f (𝓝[s] a) (𝓝 b) → is_glb (f '' s) b := -@is_lub.is_lub_of_tendsto (order_dual α) (order_dual γ) _ _ _ _ _ _ f s a b hf.dual +@is_lub.is_lub_of_tendsto αᵒᵈ γᵒᵈ _ _ _ _ _ _ f s a b hf.dual lemma is_lub.mem_lower_bounds_of_tendsto [preorder γ] [topological_space γ] [order_closed_topology γ] {f : α → γ} {s : set α} {a : α} {b : γ} (hf : antitone_on f s) (ha : is_lub s a) (hb : tendsto f (𝓝[s] a) (𝓝 b)) : b ∈ lower_bounds (f '' s) := -@is_lub.mem_upper_bounds_of_tendsto α (order_dual γ) _ _ _ _ _ _ _ _ _ _ hf ha hb +@is_lub.mem_upper_bounds_of_tendsto α γᵒᵈ _ _ _ _ _ _ _ _ _ _ hf ha hb lemma is_lub.is_glb_of_tendsto [preorder γ] [topological_space γ] [order_closed_topology γ] : ∀ {f : α → γ} {s : set α} {a : α} {b : γ}, (antitone_on f s) → is_lub s a → s.nonempty → tendsto f (𝓝[s] a) (𝓝 b) → is_glb (f '' s) b := -@is_lub.is_lub_of_tendsto α (order_dual γ) _ _ _ _ _ _ +@is_lub.is_lub_of_tendsto α γᵒᵈ _ _ _ _ _ _ lemma is_glb.mem_upper_bounds_of_tendsto [preorder γ] [topological_space γ] [order_closed_topology γ] {f : α → γ} {s : set α} {a : α} {b : γ} (hf : antitone_on f s) (ha : is_glb s a) (hb : tendsto f (𝓝[s] a) (𝓝 b)) : b ∈ upper_bounds (f '' s) := -@is_glb.mem_lower_bounds_of_tendsto α (order_dual γ) _ _ _ _ _ _ _ _ _ _ hf ha hb +@is_glb.mem_lower_bounds_of_tendsto α γᵒᵈ _ _ _ _ _ _ _ _ _ _ hf ha hb lemma is_glb.is_lub_of_tendsto [preorder γ] [topological_space γ] [order_closed_topology γ] : ∀ {f : α → γ} {s : set α} {a : α} {b : γ}, (antitone_on f s) → is_glb s a → s.nonempty → tendsto f (𝓝[s] a) (𝓝 b) → is_lub (f '' s) b := -@is_glb.is_glb_of_tendsto α (order_dual γ) _ _ _ _ _ _ +@is_glb.is_glb_of_tendsto α γᵒᵈ _ _ _ _ _ _ lemma is_lub.mem_of_is_closed {a : α} {s : set α} (ha : is_lub s a) (hs : s.nonempty) (sc : is_closed s) : a ∈ s := @@ -2191,13 +2186,13 @@ lemma is_glb.exists_seq_strict_anti_tendsto_of_not_mem {t : set α} {x : α} [is_countably_generated (𝓝 x)] (htx : is_glb t x) (not_mem : x ∉ t) (ht : t.nonempty) : ∃ u : ℕ → α, strict_anti u ∧ (∀ n, x < u n) ∧ tendsto u at_top (𝓝 x) ∧ (∀ n, u n ∈ t) := -@is_lub.exists_seq_strict_mono_tendsto_of_not_mem (order_dual α) _ _ _ t x _ htx not_mem ht +@is_lub.exists_seq_strict_mono_tendsto_of_not_mem αᵒᵈ _ _ _ t x _ htx not_mem ht lemma is_glb.exists_seq_antitone_tendsto {t : set α} {x : α} [is_countably_generated (𝓝 x)] (htx : is_glb t x) (ht : t.nonempty) : ∃ u : ℕ → α, antitone u ∧ (∀ n, x ≤ u n) ∧ tendsto u at_top (𝓝 x) ∧ (∀ n, u n ∈ t) := -@is_lub.exists_seq_monotone_tendsto (order_dual α) _ _ _ t x _ htx ht +@is_lub.exists_seq_monotone_tendsto αᵒᵈ _ _ _ t x _ htx ht lemma exists_seq_strict_anti_tendsto' [densely_ordered α] [first_countable_topology α] {x y : α} (hy : x < y) : @@ -2207,7 +2202,7 @@ by simpa only [dual_Ioo] using exists_seq_strict_mono_tendsto' (order_dual.to_du lemma exists_seq_strict_anti_tendsto [densely_ordered α] [no_max_order α] [first_countable_topology α] (x : α) : ∃ u : ℕ → α, strict_anti u ∧ (∀ n, x < u n) ∧ tendsto u at_top (𝓝 x) := -@exists_seq_strict_mono_tendsto (order_dual α) _ _ _ _ _ _ x +@exists_seq_strict_mono_tendsto αᵒᵈ _ _ _ _ _ _ x lemma exists_seq_strict_anti_strict_mono_tendsto [densely_ordered α] [first_countable_topology α] {x y : α} (h : x < y) : @@ -2224,7 +2219,7 @@ lemma exists_seq_tendsto_Inf {α : Type*} [conditionally_complete_linear_order [topological_space α] [order_topology α] [first_countable_topology α] {S : set α} (hS : S.nonempty) (hS' : bdd_below S) : ∃ (u : ℕ → α), antitone u ∧ tendsto u at_top (𝓝 (Inf S)) ∧ (∀ n, u n ∈ S) := -@exists_seq_tendsto_Sup (order_dual α) _ _ _ _ S hS hS' +@exists_seq_tendsto_Sup αᵒᵈ _ _ _ _ S hS hS' end order_topology @@ -2251,9 +2246,7 @@ closure_Ioi' nonempty_Ioi /-- The closure of the interval `(-∞, a)` is the closed interval `(-∞, a]`, unless `a` is a bottom element. -/ -lemma closure_Iio' {a : α} (h : (Iio a).nonempty) : - closure (Iio a) = Iic a := -@closure_Ioi' (order_dual α) _ _ _ _ _ h +lemma closure_Iio' (h : (Iio a).nonempty) : closure (Iio a) = Iic a := @closure_Ioi' αᵒᵈ _ _ _ _ _ h /-- The closure of the interval `(-∞, a)` is the interval `(-∞, a]`. -/ @[simp] lemma closure_Iio (a : α) [no_min_order α] : @@ -2301,7 +2294,7 @@ lemma interior_Ici [no_min_order α] {a : α} : interior (Ici a) = Ioi a := interior_Ici' nonempty_Iio @[simp] lemma interior_Iic' {a : α} (ha : (Ioi a).nonempty) : interior (Iic a) = Iio a := -@interior_Ici' (order_dual α) _ _ _ _ _ ha +@interior_Ici' αᵒᵈ _ _ _ _ _ ha lemma interior_Iic [no_max_order α] {a : α} : interior (Iic a) = Iio a := interior_Iic' nonempty_Ioi @@ -2413,7 +2406,7 @@ nhds_within_Iio_ne_bot (le_refl a) lemma filter.eventually.exists_lt [no_min_order α] {a : α} {p : α → Prop} (h : ∀ᶠ x in 𝓝 a, p x) : ∃ b < a, p b := -@filter.eventually.exists_gt (order_dual α) _ _ _ _ _ _ _ h +@filter.eventually.exists_gt αᵒᵈ _ _ _ _ _ _ _ h lemma right_nhds_within_Ico_ne_bot {a b : α} (H : a < b) : ne_bot (𝓝[Ico a b] b) := (is_lub_Ico H).nhds_within_ne_bot (nonempty_Ico.2 H) @@ -2496,7 +2489,7 @@ comap_coe_nhds_within_Ioi_of_Ioo_subset (subset.refl _) $ lemma comap_coe_Iio_nhds_within_Iio (a : α) : comap (coe : Iio a → α) (𝓝[<] a) = at_top := -@comap_coe_Ioi_nhds_within_Ioi (order_dual α) _ _ _ _ a +@comap_coe_Ioi_nhds_within_Ioi αᵒᵈ _ _ _ _ a @[simp] lemma map_coe_Ioo_at_top {a b : α} (h : a < b) : map (coe : Ioo a b → α) at_top = 𝓝[<] b := @@ -2512,7 +2505,7 @@ map_coe_at_bot_of_Ioo_subset (subset.refl _) $ λ b hb, ⟨b, hb, Ioo_subset_Ioi @[simp] lemma map_coe_Iio_at_top (a : α) : map (coe : Iio a → α) at_top = 𝓝[<] a := -@map_coe_Ioi_at_bot (order_dual α) _ _ _ _ _ +@map_coe_Ioi_at_bot αᵒᵈ _ _ _ _ _ variables {l : filter β} {f : α → β} @@ -2667,32 +2660,28 @@ the infimum of the image of this set. -/ lemma map_Inf_of_continuous_at_of_monotone' {f : α → β} {s : set α} (Cf : continuous_at f (Inf s)) (Mf : monotone f) (hs : s.nonempty) : f (Inf s) = Inf (f '' s) := -@map_Sup_of_continuous_at_of_monotone' (order_dual α) (order_dual β) _ _ _ _ _ _ f s Cf - Mf.dual hs +@map_Sup_of_continuous_at_of_monotone' αᵒᵈ βᵒᵈ _ _ _ _ _ _ f s Cf Mf.dual hs /-- A monotone function `s` sending `top` to `top` and continuous at the infimum of a set sends this infimum to the infimum of the image of this set. -/ lemma map_Inf_of_continuous_at_of_monotone {f : α → β} {s : set α} (Cf : continuous_at f (Inf s)) (Mf : monotone f) (ftop : f ⊤ = ⊤) : f (Inf s) = Inf (f '' s) := -@map_Sup_of_continuous_at_of_monotone (order_dual α) (order_dual β) _ _ _ _ _ _ f s Cf - Mf.dual ftop +@map_Sup_of_continuous_at_of_monotone αᵒᵈ βᵒᵈ _ _ _ _ _ _ f s Cf Mf.dual ftop /-- A monotone function continuous at the indexed infimum over a nonempty `Sort` sends this indexed infimum to the indexed infimum of the composition. -/ lemma map_infi_of_continuous_at_of_monotone' {ι : Sort*} [nonempty ι] {f : α → β} {g : ι → α} (Cf : continuous_at f (infi g)) (Mf : monotone f) : f (⨅ i, g i) = ⨅ i, f (g i) := -@map_supr_of_continuous_at_of_monotone' (order_dual α) (order_dual β) _ _ _ _ _ _ ι _ f g Cf - Mf.dual +@map_supr_of_continuous_at_of_monotone' αᵒᵈ βᵒᵈ _ _ _ _ _ _ ι _ f g Cf Mf.dual /-- If a monotone function sending `top` to `top` is continuous at the indexed infimum over a `Sort`, then it sends this indexed infimum to the indexed infimum of the composition. -/ lemma map_infi_of_continuous_at_of_monotone {ι : Sort*} {f : α → β} {g : ι → α} (Cf : continuous_at f (infi g)) (Mf : monotone f) (ftop : f ⊤ = ⊤) : f (infi g) = infi (f ∘ g) := -@map_supr_of_continuous_at_of_monotone (order_dual α) (order_dual β) _ _ _ _ _ _ ι f g Cf - Mf.dual ftop +@map_supr_of_continuous_at_of_monotone αᵒᵈ βᵒᵈ _ _ _ _ _ _ ι f g Cf Mf.dual ftop end complete_linear_order @@ -2738,16 +2727,14 @@ then it sends this infimum to the infimum of the image of `s`. -/ lemma map_cInf_of_continuous_at_of_monotone {f : α → β} {s : set α} (Cf : continuous_at f (Inf s)) (Mf : monotone f) (ne : s.nonempty) (H : bdd_below s) : f (Inf s) = Inf (f '' s) := -@map_cSup_of_continuous_at_of_monotone (order_dual α) (order_dual β) _ _ _ _ _ _ f s Cf - Mf.dual ne H +@map_cSup_of_continuous_at_of_monotone αᵒᵈ βᵒᵈ _ _ _ _ _ _ f s Cf Mf.dual ne H /-- A continuous monotone function sends indexed infimum to indexed infimum in conditionally complete linear order, under a boundedness assumption. -/ lemma map_cinfi_of_continuous_at_of_monotone {f : α → β} {g : γ → α} (Cf : continuous_at f (⨅ i, g i)) (Mf : monotone f) (H : bdd_below (range g)) : f (⨅ i, g i) = ⨅ i, f (g i) := -@map_csupr_of_continuous_at_of_monotone (order_dual α) (order_dual β) _ _ _ _ _ _ _ _ _ _ - Cf Mf.dual H +@map_csupr_of_continuous_at_of_monotone αᵒᵈ βᵒᵈ _ _ _ _ _ _ _ _ _ _ Cf Mf.dual H /-- A monotone map has a limit to the left of any point `x`, equal to `Sup (f '' (Iio x))`. -/ lemma monotone.tendsto_nhds_within_Iio @@ -2772,7 +2759,7 @@ lemma monotone.tendsto_nhds_within_Ioi {α : Type*} [linear_order α] [topological_space α] [order_topology α] {f : α → β} (Mf : monotone f) (x : α) : tendsto f (𝓝[>] x) (𝓝 (Inf (f '' (Ioi x)))) := -@monotone.tendsto_nhds_within_Iio (order_dual β) _ _ _ (order_dual α) _ _ _ f Mf.dual x +@monotone.tendsto_nhds_within_Iio βᵒᵈ _ _ _ αᵒᵈ _ _ _ f Mf.dual x end conditionally_complete_linear_order diff --git a/src/topology/algebra/order/compact.lean b/src/topology/algebra/order/compact.lean index 995c7f12478dd..469fa59f229be 100644 --- a/src/topology/algebra/order/compact.lean +++ b/src/topology/algebra/order/compact.lean @@ -142,9 +142,8 @@ lemma is_compact.Inf_mem {s : set α} (hs : is_compact s) (ne_s : s.nonempty) : Inf s ∈ s := hs.is_closed.cInf_mem ne_s hs.bdd_below -lemma is_compact.Sup_mem {s : set α} (hs : is_compact s) (ne_s : s.nonempty) : - Sup s ∈ s := -@is_compact.Inf_mem (order_dual α) _ _ _ _ hs ne_s +lemma is_compact.Sup_mem {s : set α} (hs : is_compact s) (ne_s : s.nonempty) : Sup s ∈ s := +@is_compact.Inf_mem αᵒᵈ _ _ _ _ hs ne_s lemma is_compact.is_glb_Inf {s : set α} (hs : is_compact s) (ne_s : s.nonempty) : is_glb s (Inf s) := @@ -152,7 +151,7 @@ is_glb_cInf ne_s hs.bdd_below lemma is_compact.is_lub_Sup {s : set α} (hs : is_compact s) (ne_s : s.nonempty) : is_lub s (Sup s) := -@is_compact.is_glb_Inf (order_dual α) _ _ _ _ hs ne_s +@is_compact.is_glb_Inf αᵒᵈ _ _ _ _ hs ne_s lemma is_compact.is_least_Inf {s : set α} (hs : is_compact s) (ne_s : s.nonempty) : is_least s (Inf s) := @@ -160,7 +159,7 @@ lemma is_compact.is_least_Inf {s : set α} (hs : is_compact s) (ne_s : s.nonempt lemma is_compact.is_greatest_Sup {s : set α} (hs : is_compact s) (ne_s : s.nonempty) : is_greatest s (Sup s) := -@is_compact.is_least_Inf (order_dual α) _ _ _ _ hs ne_s +@is_compact.is_least_Inf αᵒᵈ _ _ _ _ hs ne_s lemma is_compact.exists_is_least {s : set α} (hs : is_compact s) (ne_s : s.nonempty) : ∃ x, is_least s x := @@ -188,7 +187,7 @@ in ⟨x, hxs, hx.symm, λ y hy, lemma is_compact.exists_Sup_image_eq_and_ge {s : set β} (hs : is_compact s) (ne_s : s.nonempty) {f : β → α} (hf : continuous_on f s) : ∃ x ∈ s, Sup (f '' s) = f x ∧ ∀ y ∈ s, f y ≤ f x := -@is_compact.exists_Inf_image_eq_and_le (order_dual α) _ _ _ _ _ _ hs ne_s _ hf +@is_compact.exists_Inf_image_eq_and_le αᵒᵈ _ _ _ _ _ _ hs ne_s _ hf lemma is_compact.exists_Inf_image_eq {s : set β} (hs : is_compact s) (ne_s : s.nonempty) {f : β → α} (hf : continuous_on f s) : @@ -198,7 +197,7 @@ let ⟨x, hxs, hx, _⟩ := hs.exists_Inf_image_eq_and_le ne_s hf in ⟨x, hxs, h lemma is_compact.exists_Sup_image_eq : ∀ {s : set β}, is_compact s → s.nonempty → ∀ {f : β → α}, continuous_on f s → ∃ x ∈ s, Sup (f '' s) = f x := -@is_compact.exists_Inf_image_eq (order_dual α) _ _ _ _ _ +@is_compact.exists_Inf_image_eq αᵒᵈ _ _ _ _ _ lemma eq_Icc_of_connected_compact {s : set α} (h₁ : is_connected s) (h₂ : is_compact s) : s = Icc (Inf s) (Sup s) := @@ -222,7 +221,7 @@ end lemma is_compact.exists_forall_ge : ∀ {s : set β}, is_compact s → s.nonempty → ∀ {f : β → α}, continuous_on f s → ∃x∈s, ∀y∈s, f y ≤ f x := -@is_compact.exists_forall_le (order_dual α) _ _ _ _ _ +@is_compact.exists_forall_le αᵒᵈ _ _ _ _ _ /-- The **extreme value theorem**: if a function `f` is continuous on a closed set `s` and it is larger than a value in its image away from compact sets, then it has a minimum on this set. -/ @@ -244,7 +243,7 @@ smaller than a value in its image away from compact sets, then it has a maximum lemma continuous_on.exists_forall_ge' {s : set β} {f : β → α} (hf : continuous_on f s) (hsc : is_closed s) {x₀ : β} (h₀ : x₀ ∈ s) (hc : ∀ᶠ x in cocompact β ⊓ 𝓟 s, f x ≤ f x₀) : ∃ x ∈ s, ∀ y ∈ s, f y ≤ f x := -@continuous_on.exists_forall_le' (order_dual α) _ _ _ _ _ _ _ hf hsc _ h₀ hc +@continuous_on.exists_forall_le' αᵒᵈ _ _ _ _ _ _ _ hf hsc _ h₀ hc /-- The **extreme value theorem**: if a continuous function `f` is larger than a value in its range away from compact sets, then it has a global minimum. -/ @@ -258,7 +257,7 @@ in ⟨x, λ y, hx y (mem_univ y)⟩ away from compact sets, then it has a global maximum. -/ lemma _root_.continuous.exists_forall_ge' {f : β → α} (hf : continuous f) (x₀ : β) (h : ∀ᶠ x in cocompact β, f x ≤ f x₀) : ∃ (x : β), ∀ (y : β), f y ≤ f x := -@continuous.exists_forall_le' (order_dual α) _ _ _ _ _ _ hf x₀ h +@continuous.exists_forall_le' αᵒᵈ _ _ _ _ _ _ hf x₀ h /-- The **extreme value theorem**: if a continuous function `f` tends to infinity away from compact sets, then it has a global minimum. -/ @@ -272,7 +271,7 @@ compact sets, then it has a global maximum. -/ lemma continuous.exists_forall_ge [nonempty β] {f : β → α} (hf : continuous f) (hlim : tendsto f (cocompact β) at_bot) : ∃ x, ∀ y, f y ≤ f x := -@continuous.exists_forall_le (order_dual α) _ _ _ _ _ _ _ hf hlim +@continuous.exists_forall_le αᵒᵈ _ _ _ _ _ _ _ hf hlim lemma is_compact.Sup_lt_iff_of_continuous {f : β → α} {K : set β} (hK : is_compact K) (h0K : K.nonempty) (hf : continuous_on f K) (y : α) : @@ -289,7 +288,7 @@ lemma is_compact.lt_Inf_iff_of_continuous {α β : Type*} [order_topology α] [topological_space β] {f : β → α} {K : set β} (hK : is_compact K) (h0K : K.nonempty) (hf : continuous_on f K) (y : α) : y < Inf (f '' K) ↔ ∀ x ∈ K, y < f x := -@is_compact.Sup_lt_iff_of_continuous (order_dual α) β _ _ _ _ _ _ hK h0K hf y +@is_compact.Sup_lt_iff_of_continuous αᵒᵈ β _ _ _ _ _ _ hK h0K hf y /-- A continuous function with compact support has a global minimum. -/ @[to_additive] @@ -307,7 +306,7 @@ end lemma continuous.exists_forall_ge_of_has_compact_mul_support [nonempty β] [has_one α] {f : β → α} (hf : continuous f) (h : has_compact_mul_support f) : ∃ (x : β), ∀ (y : β), f y ≤ f x := -@continuous.exists_forall_le_of_has_compact_mul_support (order_dual α) _ _ _ _ _ _ _ _ hf h +@continuous.exists_forall_le_of_has_compact_mul_support αᵒᵈ _ _ _ _ _ _ _ _ hf h lemma is_compact.continuous_Sup {f : γ → β → α} {K : set β} (hK : is_compact K) (hf : continuous ↿f) : @@ -339,7 +338,7 @@ end lemma is_compact.continuous_Inf {f : γ → β → α} {K : set β} (hK : is_compact K) (hf : continuous ↿f) : continuous (λ x, Inf (f x '' K)) := -@is_compact.continuous_Sup (order_dual α) β γ _ _ _ _ _ _ _ hK hf +@is_compact.continuous_Sup αᵒᵈ β γ _ _ _ _ _ _ _ hK hf namespace continuous_on /-! diff --git a/src/topology/algebra/order/extend_from.lean b/src/topology/algebra/order/extend_from.lean index 9078a10044091..9578d72199b64 100644 --- a/src/topology/algebra/order/extend_from.lean +++ b/src/topology/algebra/order/extend_from.lean @@ -74,7 +74,7 @@ lemma continuous_on_Ioc_extend_from_Ioo [topological_space α] (hb : tendsto f (𝓝[<] b) (𝓝 lb)) : continuous_on (extend_from (Ioo a b) f) (Ioc a b) := begin - have := @continuous_on_Ico_extend_from_Ioo (order_dual α) _ _ _ _ _ _ _ f _ _ _ hab, + have := @continuous_on_Ico_extend_from_Ioo αᵒᵈ _ _ _ _ _ _ _ f _ _ _ hab, erw [dual_Ico, dual_Ioi, dual_Ioo] at this, exact this hf hb end diff --git a/src/topology/algebra/order/intermediate_value.lean b/src/topology/algebra/order/intermediate_value.lean index 18bce4b1a6321..83e6f74b73028 100644 --- a/src/topology/algebra/order/intermediate_value.lean +++ b/src/topology/algebra/order/intermediate_value.lean @@ -267,7 +267,7 @@ end lemma is_preconnected.Iio_cSup_subset {s : set α} (hs : is_preconnected s) (hb : ¬bdd_below s) (ha : bdd_above s) : Iio (Sup s) ⊆ s := -@is_preconnected.Ioi_cInf_subset (order_dual α) _ _ _ s hs ha hb +@is_preconnected.Ioi_cInf_subset αᵒᵈ _ _ _ s hs ha hb /-- A preconnected set in a conditionally complete linear order is either one of the intervals `[Inf s, Sup s]`, `[Inf s, Sup s)`, `(Inf s, Sup s]`, `(Inf s, Sup s)`, `[Inf s, +∞)`, @@ -547,7 +547,7 @@ lemma continuous.surjective {f : α → δ} (hf : continuous f) (h_top : tendsto lemma continuous.surjective' {f : α → δ} (hf : continuous f) (h_top : tendsto f at_bot at_top) (h_bot : tendsto f at_top at_bot) : function.surjective f := -@continuous.surjective (order_dual α) _ _ _ _ _ _ _ _ _ hf h_top h_bot +@continuous.surjective αᵒᵈ _ _ _ _ _ _ _ _ _ hf h_top h_bot /-- If a function `f : α → β` is continuous on a nonempty interval `s`, its restriction to `s` tends to `at_bot : filter β` along `at_bot : filter ↥s` and tends to `at_top : filter β` along @@ -569,4 +569,4 @@ lemma continuous_on.surj_on_of_tendsto' {f : α → δ} {s : set α} [ord_connec (hs : s.nonempty) (hf : continuous_on f s) (hbot : tendsto (λ x : s, f x) at_bot at_top) (htop : tendsto (λ x : s, f x) at_top at_bot) : surj_on f s univ := -@continuous_on.surj_on_of_tendsto α _ _ _ _ (order_dual δ) _ _ _ _ _ _ hs hf hbot htop +@continuous_on.surj_on_of_tendsto α _ _ _ _ δᵒᵈ _ _ _ _ _ _ hs hf hbot htop diff --git a/src/topology/algebra/order/left_right.lean b/src/topology/algebra/order/left_right.lean index 5540e488258f4..21e4190f327e4 100644 --- a/src/topology/algebra/order/left_right.lean +++ b/src/topology/algebra/order/left_right.lean @@ -35,7 +35,7 @@ by simp only [← Ici_diff_left, continuous_within_at_diff_self] lemma continuous_within_at_Iio_iff_Iic {a : α} {f : α → β} : continuous_within_at f (Iio a) a ↔ continuous_within_at f (Iic a) a := -@continuous_within_at_Ioi_iff_Ici (order_dual α) _ ‹topological_space α› _ _ _ f +@continuous_within_at_Ioi_iff_Ici αᵒᵈ _ ‹topological_space α› _ _ _ f end partial_order diff --git a/src/topology/algebra/order/liminf_limsup.lean b/src/topology/algebra/order/liminf_limsup.lean index e37326d87f201..30d7484a0e540 100644 --- a/src/topology/algebra/order/liminf_limsup.lean +++ b/src/topology/algebra/order/liminf_limsup.lean @@ -51,8 +51,7 @@ end order_closed_topology section order_closed_topology variables [semilattice_inf α] [topological_space α] [order_topology α] -lemma is_bounded_ge_nhds (a : α) : (𝓝 a).is_bounded (≥) := -@is_bounded_le_nhds (order_dual α) _ _ _ a +lemma is_bounded_ge_nhds (a : α) : (𝓝 a).is_bounded (≥) := @is_bounded_le_nhds αᵒᵈ _ _ _ a lemma filter.tendsto.is_bounded_under_ge {f : filter β} {u : β → α} {a : α} (h : tendsto u f (𝓝 a)) : f.is_bounded_under (≥) u := @@ -85,7 +84,7 @@ mem_of_superset h $ assume a hac, lt_of_le_of_lt hac hcb theorem gt_mem_sets_of_Liminf_gt : ∀ {f : filter α} {b}, f.is_bounded (≥) → b < f.Liminf → ∀ᶠ a in f, b < a := -@lt_mem_sets_of_Limsup_lt (order_dual α) _ +@lt_mem_sets_of_Limsup_lt αᵒᵈ _ variables [topological_space α] [order_topology α] @@ -107,8 +106,7 @@ cInf_eq_of_forall_ge_of_forall_gt_exists_lt (is_bounded_le_nhds a) | or.inr ⟨_, h⟩ := ⟨a, (𝓝 a).sets_of_superset (gt_mem_nhds hba) h, hba⟩ end) -theorem Liminf_nhds : ∀ (a : α), Liminf (𝓝 a) = a := -@Limsup_nhds (order_dual α) _ _ _ +theorem Liminf_nhds : ∀ (a : α), Liminf (𝓝 a) = a := @Limsup_nhds αᵒᵈ _ _ _ /-- If a filter is converging, its limsup coincides with its limit. -/ theorem Liminf_eq_of_le_nhds {f : filter α} {a : α} [ne_bot f] (h : f ≤ 𝓝 a) : f.Liminf = a := @@ -125,7 +123,7 @@ le_antisymm /-- If a filter is converging, its liminf coincides with its limit. -/ theorem Limsup_eq_of_le_nhds : ∀ {f : filter α} {a : α} [ne_bot f], f ≤ 𝓝 a → f.Limsup = a := -@Liminf_eq_of_le_nhds (order_dual α) _ _ _ +@Liminf_eq_of_le_nhds αᵒᵈ _ _ _ /-- If a function has a limit, then its limsup coincides with its limit. -/ theorem filter.tendsto.limsup_eq {f : filter β} {u : β → α} {a : α} [ne_bot f] diff --git a/src/topology/algebra/order/monotone_continuity.lean b/src/topology/algebra/order/monotone_continuity.lean index 9a199fdabf655..957caae659ca6 100644 --- a/src/topology/algebra/order/monotone_continuity.lean +++ b/src/topology/algebra/order/monotone_continuity.lean @@ -152,8 +152,7 @@ lemma continuous_at_left_of_monotone_on_of_exists_between {f : α → β} {s : s (hf : monotone_on f s) (hs : s ∈ 𝓝[≤] a) (hfs : ∀ b < f a, ∃ c ∈ s, f c ∈ Ioo b (f a)) : continuous_within_at f (Iic a) a := -@continuous_at_right_of_monotone_on_of_exists_between (order_dual α) (order_dual β) _ _ _ _ _ _ - f s a hf.dual hs $ +@continuous_at_right_of_monotone_on_of_exists_between αᵒᵈ βᵒᵈ _ _ _ _ _ _ f s a hf.dual hs $ λ b hb, let ⟨c, hcs, hcb, hca⟩ := hfs b hb in ⟨c, hcs, hca, hcb⟩ /-- If a function `f` with a densely ordered codomain is monotone on a left neighborhood of `a` and @@ -163,8 +162,8 @@ lemma continuous_at_left_of_monotone_on_of_closure_image_mem_nhds_within [densel {f : α → β} {s : set α} {a : α} (hf : monotone_on f s) (hs : s ∈ 𝓝[≤] a) (hfs : closure (f '' s) ∈ 𝓝[≤] (f a)) : continuous_within_at f (Iic a) a := -@continuous_at_right_of_monotone_on_of_closure_image_mem_nhds_within (order_dual α) (order_dual β) - _ _ _ _ _ _ _ f s a hf.dual hs hfs +@continuous_at_right_of_monotone_on_of_closure_image_mem_nhds_within αᵒᵈ βᵒᵈ _ _ _ _ _ _ _ + f s a hf.dual hs hfs /-- If a function `f` with a densely ordered codomain is monotone on a left neighborhood of `a` and the image of this neighborhood under `f` is a left neighborhood of `f a`, then `f` is continuous at diff --git a/src/topology/algebra/order/monotone_convergence.lean b/src/topology/algebra/order/monotone_convergence.lean index 1b72764230fe6..fa813e949b10c 100644 --- a/src/topology/algebra/order/monotone_convergence.lean +++ b/src/topology/algebra/order/monotone_convergence.lean @@ -49,11 +49,11 @@ class Inf_convergence_class (α : Type*) [preorder α] [topological_space α] : (tendsto_coe_at_bot_is_glb : ∀ (a : α) (s : set α), is_glb s a → tendsto (coe : s → α) at_bot (𝓝 a)) instance order_dual.Sup_convergence_class [preorder α] [topological_space α] - [Inf_convergence_class α] : Sup_convergence_class (order_dual α) := + [Inf_convergence_class α] : Sup_convergence_class αᵒᵈ := ⟨‹Inf_convergence_class α›.1⟩ instance order_dual.Inf_convergence_class [preorder α] [topological_space α] - [Sup_convergence_class α] : Inf_convergence_class (order_dual α) := + [Sup_convergence_class α] : Inf_convergence_class αᵒᵈ := ⟨‹Sup_convergence_class α›.1⟩ @[priority 100] -- see Note [lower instance priority] @@ -70,7 +70,7 @@ end @[priority 100] -- see Note [lower instance priority] instance linear_order.Inf_convergence_class [topological_space α] [linear_order α] [order_topology α] : Inf_convergence_class α := -show Inf_convergence_class (order_dual $ order_dual α), from order_dual.Inf_convergence_class +show Inf_convergence_class αᵒᵈᵒᵈ, from order_dual.Inf_convergence_class section @@ -88,9 +88,9 @@ begin exact h_mono.range_factorization.tendsto_at_top_at_top (λ b, b.2.imp $ λ a ha, ha.ge) end -lemma tendsto_at_bot_is_lub (h_anti : antitone f) - (ha : is_lub (set.range f) a) : tendsto f at_bot (𝓝 a) := -@tendsto_at_top_is_lub α (order_dual ι) _ _ _ _ f a h_anti.dual ha +lemma tendsto_at_bot_is_lub (h_anti : antitone f) (ha : is_lub (set.range f) a) : + tendsto f at_bot (𝓝 a) := +by convert tendsto_at_top_is_lub h_anti.dual_left ha end is_lub @@ -100,12 +100,11 @@ variables [preorder α] [Inf_convergence_class α] {f : ι → α} {a : α} lemma tendsto_at_bot_is_glb (h_mono : monotone f) (ha : is_glb (set.range f) a) : tendsto f at_bot (𝓝 a) := -@tendsto_at_top_is_lub (order_dual α) (order_dual ι) _ _ _ _ f a h_mono.dual ha +by convert tendsto_at_top_is_lub h_mono.dual ha.dual -lemma tendsto_at_top_is_glb (h_anti : antitone f) - (ha : is_glb (set.range f) a) : +lemma tendsto_at_top_is_glb (h_anti : antitone f) (ha : is_glb (set.range f) a) : tendsto f at_top (𝓝 a) := -@tendsto_at_top_is_lub (order_dual α) ι _ _ _ _ f a h_anti ha +by convert tendsto_at_bot_is_lub h_anti.dual ha.dual end is_glb @@ -120,10 +119,9 @@ begin exacts [tendsto_of_is_empty, tendsto_at_top_is_lub h_mono (is_lub_csupr hbdd)] end -lemma tendsto_at_bot_csupr (h_anti : antitone f) - (hbdd : bdd_above $ range f) : - tendsto f at_bot (𝓝 (⨆i, f i)) := -@tendsto_at_top_csupr α (order_dual ι) _ _ _ _ _ h_anti.dual hbdd +lemma tendsto_at_bot_csupr (h_anti : antitone f) (hbdd : bdd_above $ range f) : + tendsto f at_bot (𝓝 (⨆ i, f i)) := +by convert tendsto_at_top_csupr h_anti.dual hbdd.dual end csupr @@ -132,13 +130,12 @@ section cinfi variables [conditionally_complete_lattice α] [Inf_convergence_class α] {f : ι → α} {a : α} lemma tendsto_at_bot_cinfi (h_mono : monotone f) (hbdd : bdd_below $ range f) : - tendsto f at_bot (𝓝 (⨅i, f i)) := -@tendsto_at_top_csupr (order_dual α) (order_dual ι) _ _ _ _ _ h_mono.dual hbdd + tendsto f at_bot (𝓝 (⨅ i, f i)) := +by convert tendsto_at_top_csupr h_mono.dual hbdd.dual -lemma tendsto_at_top_cinfi (h_anti : antitone f) - (hbdd : bdd_below $ range f) : - tendsto f at_top (𝓝 (⨅i, f i)) := -@tendsto_at_top_csupr (order_dual α) ι _ _ _ _ _ h_anti hbdd +lemma tendsto_at_top_cinfi (h_anti : antitone f) (hbdd : bdd_below $ range f) : + tendsto f at_top (𝓝 (⨅ i, f i)) := +by convert tendsto_at_bot_csupr h_anti.dual hbdd.dual end cinfi @@ -186,8 +183,7 @@ end instance [preorder α] [preorder β] [topological_space α] [topological_space β] [Inf_convergence_class α] [Inf_convergence_class β] : Inf_convergence_class (α × β) := -show Inf_convergence_class (order_dual $ (order_dual α × order_dual β)), - from order_dual.Inf_convergence_class +show Inf_convergence_class (αᵒᵈ × βᵒᵈ)ᵒᵈ, from order_dual.Inf_convergence_class instance {ι : Type*} {α : ι → Type*} [Π i, preorder (α i)] [Π i, topological_space (α i)] [Π i, Sup_convergence_class (α i)] : Sup_convergence_class (Π i, α i) := @@ -199,8 +195,7 @@ end instance {ι : Type*} {α : ι → Type*} [Π i, preorder (α i)] [Π i, topological_space (α i)] [Π i, Inf_convergence_class (α i)] : Inf_convergence_class (Π i, α i) := -show Inf_convergence_class (order_dual $ Π i, order_dual (α i)), - from order_dual.Inf_convergence_class +show Inf_convergence_class (Π i, (α i)ᵒᵈ)ᵒᵈ, from order_dual.Inf_convergence_class instance pi.Sup_convergence_class' {ι : Type*} [preorder α] [topological_space α] [Sup_convergence_class α] : Sup_convergence_class (ι → α) := @@ -279,19 +274,19 @@ lemma is_glb_of_tendsto_at_bot [topological_space α] [preorder α] [order_close [nonempty β] [semilattice_inf β] {f : β → α} {a : α} (hf : monotone f) (ha : tendsto f at_bot (𝓝 a)) : is_glb (set.range f) a := -@is_lub_of_tendsto_at_top (order_dual α) (order_dual β) _ _ _ _ _ _ _ hf.dual ha +@is_lub_of_tendsto_at_top αᵒᵈ βᵒᵈ _ _ _ _ _ _ _ hf.dual ha lemma is_lub_of_tendsto_at_bot [topological_space α] [preorder α] [order_closed_topology α] [nonempty β] [semilattice_inf β] {f : β → α} {a : α} (hf : antitone f) (ha : tendsto f at_bot (𝓝 a)) : is_lub (set.range f) a := -@is_lub_of_tendsto_at_top α (order_dual β) _ _ _ _ _ _ _ hf.dual_left ha +@is_lub_of_tendsto_at_top α βᵒᵈ _ _ _ _ _ _ _ hf.dual_left ha lemma is_glb_of_tendsto_at_top [topological_space α] [preorder α] [order_closed_topology α] [nonempty β] [semilattice_sup β] {f : β → α} {a : α} (hf : antitone f) (ha : tendsto f at_top (𝓝 a)) : is_glb (set.range f) a := -@is_glb_of_tendsto_at_bot α (order_dual β) _ _ _ _ _ _ _ hf.dual_left ha +@is_glb_of_tendsto_at_bot α βᵒᵈ _ _ _ _ _ _ _ hf.dual_left ha lemma supr_eq_of_tendsto {α β} [topological_space α] [complete_linear_order α] [order_topology α] [nonempty β] [semilattice_sup β] {f : β → α} {a : α} (hf : monotone f) : diff --git a/src/topology/continuous_function/ordered.lean b/src/topology/continuous_function/ordered.lean index 6961a685200de..ce6e88b9f1c1a 100644 --- a/src/topology/continuous_function/ordered.lean +++ b/src/topology/continuous_function/ordered.lean @@ -110,12 +110,12 @@ variables [linear_order γ] [order_closed_topology γ] lemma inf'_apply {ι : Type*} {s : finset ι} (H : s.nonempty) (f : ι → C(β, γ)) (b : β) : s.inf' H f b = s.inf' H (λ a, f a b) := -@sup'_apply _ (order_dual γ) _ _ _ _ _ _ H f b +@sup'_apply _ γᵒᵈ _ _ _ _ _ _ H f b @[simp, norm_cast] lemma inf'_coe {ι : Type*} {s : finset ι} (H : s.nonempty) (f : ι → C(β, γ)) : ((s.inf' H f : C(β, γ)) : ι → β) = s.inf' H (λ a, (f a : β → γ)) := -@sup'_coe _ (order_dual γ) _ _ _ _ _ _ H f +@sup'_coe _ γᵒᵈ _ _ _ _ _ _ H f end inf' diff --git a/src/topology/local_extr.lean b/src/topology/local_extr.lean index bdbbe6da9b0e8..ed73d7aeb1c31 100644 --- a/src/topology/local_extr.lean +++ b/src/topology/local_extr.lean @@ -143,7 +143,7 @@ let ⟨y, hy⟩ := (this.and self_mem_nhds_within).exists in hy.1.not_lt hy.2 lemma is_local_max_on.not_nhds_le_map [topological_space β] (hf : is_local_max_on f s a) [ne_bot (𝓝[>] (f a))] : ¬𝓝 (f a) ≤ map f (𝓝[s] a) := -@is_local_min_on.not_nhds_le_map α (order_dual β) _ _ _ _ _ ‹_› hf ‹_› +@is_local_min_on.not_nhds_le_map α βᵒᵈ _ _ _ _ _ ‹_› hf ‹_› lemma is_local_extr_on.not_nhds_le_map [topological_space β] (hf : is_local_extr_on f s a) [ne_bot (𝓝[<] (f a))] [ne_bot (𝓝[>] (f a))] : diff --git a/src/topology/order.lean b/src/topology/order.lean index ec2b2d47ce9f7..ede00f242234c 100644 --- a/src/topology/order.lean +++ b/src/topology/order.lean @@ -808,56 +808,54 @@ variables {α : Type u} {ι : Sort v} lemma generate_from_union (a₁ a₂ : set (set α)) : topological_space.generate_from (a₁ ∪ a₂) = topological_space.generate_from a₁ ⊓ topological_space.generate_from a₂ := -@galois_connection.l_sup _ (order_dual (topological_space α)) a₁ a₂ _ _ _ _ +@galois_connection.l_sup _ (topological_space α)ᵒᵈ a₁ a₂ _ _ _ _ (λ g t, generate_from_le_iff_subset_is_open) lemma set_of_is_open_sup (t₁ t₂ : topological_space α) : {s | (t₁ ⊔ t₂).is_open s} = {s | t₁.is_open s} ∩ {s | t₂.is_open s} := -@galois_connection.u_inf _ (order_dual (topological_space α)) t₁ t₂ _ _ _ _ +@galois_connection.u_inf _ (topological_space α)ᵒᵈ t₁ t₂ _ _ _ _ (λ g t, generate_from_le_iff_subset_is_open) lemma generate_from_Union {f : ι → set (set α)} : topological_space.generate_from (⋃ i, f i) = (⨅ i, topological_space.generate_from (f i)) := -@galois_connection.l_supr _ (order_dual (topological_space α)) _ _ _ _ _ +@galois_connection.l_supr _ (topological_space α)ᵒᵈ _ _ _ _ _ (λ g t, generate_from_le_iff_subset_is_open) f lemma set_of_is_open_supr {t : ι → topological_space α} : {s | (⨆ i, t i).is_open s} = ⋂ i, {s | (t i).is_open s} := -@galois_connection.u_infi _ (order_dual (topological_space α)) _ _ _ _ _ +@galois_connection.u_infi _ (topological_space α)ᵒᵈ _ _ _ _ _ (λ g t, generate_from_le_iff_subset_is_open) t lemma generate_from_sUnion {S : set (set (set α))} : topological_space.generate_from (⋃₀ S) = (⨅ s ∈ S, topological_space.generate_from s) := -@galois_connection.l_Sup _ (order_dual (topological_space α)) _ _ _ _ +@galois_connection.l_Sup _ (topological_space α)ᵒᵈ _ _ _ _ (λ g t, generate_from_le_iff_subset_is_open) S lemma set_of_is_open_Sup {T : set (topological_space α)} : {s | (Sup T).is_open s} = ⋂ t ∈ T, {s | (t : topological_space α).is_open s} := -@galois_connection.u_Inf _ (order_dual (topological_space α)) _ _ _ _ +@galois_connection.u_Inf _ (topological_space α)ᵒᵈ _ _ _ _ (λ g t, generate_from_le_iff_subset_is_open) T lemma generate_from_union_is_open (a b : topological_space α) : topological_space.generate_from ({s | a.is_open s} ∪ {s | b.is_open s}) = a ⊓ b := -@galois_insertion.l_sup_u _ (order_dual (topological_space α)) _ _ _ _ (gi_generate_from α) a b +@galois_insertion.l_sup_u _ (topological_space α)ᵒᵈ _ _ _ _ (gi_generate_from α) a b lemma generate_from_Union_is_open (f : ι → topological_space α) : topological_space.generate_from (⋃ i, {s | (f i).is_open s}) = ⨅ i, (f i) := -@galois_insertion.l_supr_u _ (order_dual (topological_space α)) _ _ _ _ (gi_generate_from α) _ f +@galois_insertion.l_supr_u _ (topological_space α)ᵒᵈ _ _ _ _ (gi_generate_from α) _ f lemma generate_from_inter (a b : topological_space α) : topological_space.generate_from ({s | a.is_open s} ∩ {s | b.is_open s}) = a ⊔ b := -@galois_insertion.l_inf_u _ (order_dual (topological_space α)) _ _ _ _ - (gi_generate_from α) a b +@galois_insertion.l_inf_u _ (topological_space α)ᵒᵈ _ _ _ _ (gi_generate_from α) a b lemma generate_from_Inter (f : ι → topological_space α) : topological_space.generate_from (⋂ i, {s | (f i).is_open s}) = ⨆ i, (f i) := -@galois_insertion.l_infi_u _ (order_dual (topological_space α)) _ _ _ _ (gi_generate_from α) _ f +@galois_insertion.l_infi_u _ (topological_space α)ᵒᵈ _ _ _ _ (gi_generate_from α) _ f lemma generate_from_Inter_of_generate_from_eq_self (f : ι → set (set α)) (hf : ∀ i, {s | (topological_space.generate_from (f i)).is_open s} = f i) : topological_space.generate_from (⋂ i, (f i)) = ⨆ i, topological_space.generate_from (f i) := -@galois_insertion.l_infi_of_ul_eq_self _ (order_dual (topological_space α)) _ _ _ _ - (gi_generate_from α) _ f hf +@galois_insertion.l_infi_of_ul_eq_self _ (topological_space α)ᵒᵈ _ _ _ _ (gi_generate_from α) _ f hf variables {t : ι → topological_space α} diff --git a/src/topology/order/lattice.lean b/src/topology/order/lattice.lean index 71156f5249437..e42a6d380704e 100644 --- a/src/topology/order/lattice.lean +++ b/src/topology/order/lattice.lean @@ -43,14 +43,12 @@ class has_continuous_sup (L : Type*) [topological_space L] [has_sup L] : Prop := @[priority 100] -- see Note [lower instance priority] instance order_dual.has_continuous_sup - (L : Type*) [topological_space L] [has_inf L] [has_continuous_inf L] : - has_continuous_sup (order_dual L) := + (L : Type*) [topological_space L] [has_inf L] [has_continuous_inf L] : has_continuous_sup Lᵒᵈ := { continuous_sup := @has_continuous_inf.continuous_inf L _ _ _ } @[priority 100] -- see Note [lower instance priority] instance order_dual.has_continuous_inf - (L : Type*) [topological_space L] [has_sup L] [has_continuous_sup L] : - has_continuous_inf (order_dual L) := + (L : Type*) [topological_space L] [has_sup L] [has_continuous_sup L] : has_continuous_inf Lᵒᵈ := { continuous_inf := @has_continuous_sup.continuous_sup L _ _ _ } /-- @@ -63,7 +61,7 @@ class topological_lattice (L : Type*) [topological_space L] [lattice L] @[priority 100] -- see Note [lower instance priority] instance order_dual.topological_lattice (L : Type*) [topological_space L] [lattice L] [topological_lattice L] : - topological_lattice (order_dual L) := {} + topological_lattice Lᵒᵈ := {} variables {L : Type*} [topological_space L] variables {X : Type*} [topological_space X] diff --git a/src/topology/semicontinuous.lean b/src/topology/semicontinuous.lean index d62e6cd694c96..bec1d2dd141c2 100644 --- a/src/topology/semicontinuous.lean +++ b/src/topology/semicontinuous.lean @@ -297,13 +297,12 @@ lemma continuous.comp_lower_semicontinuous lemma continuous_at.comp_lower_semicontinuous_within_at_antitone {g : γ → δ} {f : α → γ} (hg : continuous_at g (f x)) (hf : lower_semicontinuous_within_at f s x) (gmon : antitone g) : upper_semicontinuous_within_at (g ∘ f) s x := -@continuous_at.comp_lower_semicontinuous_within_at α _ x s γ _ _ _ (order_dual δ) _ _ _ - g f hg hf gmon +@continuous_at.comp_lower_semicontinuous_within_at α _ x s γ _ _ _ δᵒᵈ _ _ _ g f hg hf gmon lemma continuous_at.comp_lower_semicontinuous_at_antitone {g : γ → δ} {f : α → γ} (hg : continuous_at g (f x)) (hf : lower_semicontinuous_at f x) (gmon : antitone g) : upper_semicontinuous_at (g ∘ f) x := -@continuous_at.comp_lower_semicontinuous_at α _ x γ _ _ _ (order_dual δ) _ _ _ g f hg hf gmon +@continuous_at.comp_lower_semicontinuous_at α _ x γ _ _ _ δᵒᵈ _ _ _ g f hg hf gmon lemma continuous.comp_lower_semicontinuous_on_antitone {g : γ → δ} {f : α → γ} (hg : continuous g) (hf : lower_semicontinuous_on f s) @@ -632,7 +631,7 @@ variables [has_zero β] lemma is_open.upper_semicontinuous_indicator (hs : is_open s) (hy : y ≤ 0) : upper_semicontinuous (indicator s (λ x, y)) := -@is_open.lower_semicontinuous_indicator α _ (order_dual β) _ s y _ hs hy +@is_open.lower_semicontinuous_indicator α _ βᵒᵈ _ s y _ hs hy lemma is_open.upper_semicontinuous_on_indicator (hs : is_open s) (hy : y ≤ 0) : upper_semicontinuous_on (indicator s (λ x, y)) t := @@ -648,7 +647,7 @@ lemma is_open.upper_semicontinuous_within_at_indicator (hs : is_open s) (hy : y lemma is_closed.upper_semicontinuous_indicator (hs : is_closed s) (hy : 0 ≤ y) : upper_semicontinuous (indicator s (λ x, y)) := -@is_closed.lower_semicontinuous_indicator α _ (order_dual β) _ s y _ hs hy +@is_closed.lower_semicontinuous_indicator α _ βᵒᵈ _ s y _ hs hy lemma is_closed.upper_semicontinuous_on_indicator (hs : is_closed s) (hy : 0 ≤ y) : upper_semicontinuous_on (indicator s (λ x, y)) t := @@ -704,14 +703,12 @@ variables {δ : Type*} [linear_order δ] [topological_space δ] [order_topology lemma continuous_at.comp_upper_semicontinuous_within_at {g : γ → δ} {f : α → γ} (hg : continuous_at g (f x)) (hf : upper_semicontinuous_within_at f s x) (gmon : monotone g) : upper_semicontinuous_within_at (g ∘ f) s x := -@continuous_at.comp_lower_semicontinuous_within_at α _ x s (order_dual γ) _ _ _ - (order_dual δ) _ _ _ g f hg hf (λ x y hxy, gmon hxy) +@continuous_at.comp_lower_semicontinuous_within_at α _ x s γᵒᵈ _ _ _ δᵒᵈ _ _ _ g f hg hf gmon.dual lemma continuous_at.comp_upper_semicontinuous_at {g : γ → δ} {f : α → γ} (hg : continuous_at g (f x)) (hf : upper_semicontinuous_at f x) (gmon : monotone g) : upper_semicontinuous_at (g ∘ f) x := -@continuous_at.comp_lower_semicontinuous_at α _ x (order_dual γ) _ _ _ - (order_dual δ) _ _ _ g f hg hf (λ x y hxy, gmon hxy) +@continuous_at.comp_lower_semicontinuous_at α _ x γᵒᵈ _ _ _ δᵒᵈ _ _ _ g f hg hf gmon.dual lemma continuous.comp_upper_semicontinuous_on {g : γ → δ} {f : α → γ} (hg : continuous g) (hf : upper_semicontinuous_on f s) @@ -726,13 +723,12 @@ lemma continuous.comp_upper_semicontinuous lemma continuous_at.comp_upper_semicontinuous_within_at_antitone {g : γ → δ} {f : α → γ} (hg : continuous_at g (f x)) (hf : upper_semicontinuous_within_at f s x) (gmon : antitone g) : lower_semicontinuous_within_at (g ∘ f) s x := -@continuous_at.comp_upper_semicontinuous_within_at α _ x s γ _ _ _ (order_dual δ) _ _ _ - g f hg hf gmon +@continuous_at.comp_upper_semicontinuous_within_at α _ x s γ _ _ _ δᵒᵈ _ _ _ g f hg hf gmon lemma continuous_at.comp_upper_semicontinuous_at_antitone {g : γ → δ} {f : α → γ} (hg : continuous_at g (f x)) (hf : upper_semicontinuous_at f x) (gmon : antitone g) : lower_semicontinuous_at (g ∘ f) x := -@continuous_at.comp_upper_semicontinuous_at α _ x γ _ _ _ (order_dual δ) _ _ _ g f hg hf gmon +@continuous_at.comp_upper_semicontinuous_at α _ x γ _ _ _ δᵒᵈ _ _ _ g f hg hf gmon lemma continuous.comp_upper_semicontinuous_on_antitone {g : γ → δ} {f : α → γ} (hg : continuous g) (hf : upper_semicontinuous_on f s) @@ -759,7 +755,7 @@ lemma upper_semicontinuous_within_at.add' {f g : α → γ} (hf : upper_semicontinuous_within_at f s x) (hg : upper_semicontinuous_within_at g s x) (hcont : continuous_at (λ (p : γ × γ), p.1 + p.2) (f x, g x)) : upper_semicontinuous_within_at (λ z, f z + g z) s x := -@lower_semicontinuous_within_at.add' α _ x s (order_dual γ) _ _ _ _ _ hf hg hcont +@lower_semicontinuous_within_at.add' α _ x s γᵒᵈ _ _ _ _ _ hf hg hcont /-- The sum of two upper semicontinuous functions is upper semicontinuous. Formulated with an explicit continuity assumption on addition, for application to `ereal`. The unprimed version of @@ -825,7 +821,7 @@ hf.add' hg (λ x, continuous_add.continuous_at) lemma upper_semicontinuous_within_at_sum {f : ι → α → γ} {a : finset ι} (ha : ∀ i ∈ a, upper_semicontinuous_within_at (f i) s x) : upper_semicontinuous_within_at (λ z, (∑ i in a, f i z)) s x := -@lower_semicontinuous_within_at_sum α _ x s ι (order_dual γ) _ _ _ _ f a ha +@lower_semicontinuous_within_at_sum α _ x s ι γᵒᵈ _ _ _ _ f a ha lemma upper_semicontinuous_at_sum {f : ι → α → γ} {a : finset ι} (ha : ∀ i ∈ a, upper_semicontinuous_at (f i) x) : @@ -855,7 +851,7 @@ variables {ι : Sort*} {δ : Type*} [complete_linear_order δ] lemma upper_semicontinuous_within_at_infi {f : ι → α → δ} (h : ∀ i, upper_semicontinuous_within_at (f i) s x) : upper_semicontinuous_within_at (λ x', ⨅ i, f i x') s x := -@lower_semicontinuous_within_at_supr α _ x s ι (order_dual δ) _ f h +@lower_semicontinuous_within_at_supr α _ x s ι δᵒᵈ _ f h lemma upper_semicontinuous_within_at_binfi {p : ι → Prop} {f : Π i (h : p i), α → δ} (h : ∀ i hi, upper_semicontinuous_within_at (f i hi) s x) : @@ -865,7 +861,7 @@ upper_semicontinuous_within_at_infi $ λ i, upper_semicontinuous_within_at_infi lemma upper_semicontinuous_at_infi {f : ι → α → δ} (h : ∀ i, upper_semicontinuous_at (f i) x) : upper_semicontinuous_at (λ x', ⨅ i, f i x') x := -@lower_semicontinuous_at_supr α _ x ι (order_dual δ) _ f h +@lower_semicontinuous_at_supr α _ x ι δᵒᵈ _ f h lemma upper_semicontinuous_at_binfi {p : ι → Prop} {f : Π i (h : p i), α → δ} (h : ∀ i hi, upper_semicontinuous_at (f i hi) x) : From 53c79d561eec6b94be44bdaedeb2621361aa514f Mon Sep 17 00:00:00 2001 From: Haruhisa Enomoto <73296681+haruhisa-enomoto@users.noreply.github.com> Date: Wed, 4 May 2022 11:10:46 +0000 Subject: [PATCH 097/153] feat(linear_algebra/span): add `finite_span_is_compact_element` (#13901) This PR adds `finite_span_is_compact_element`, which extends `singleton_span_is_compact_element` to the spans of finite subsets. This will be useful e.g. when proving the existence of a maximal submodule of a finitely generated module. --- src/linear_algebra/span.lean | 16 ++++++++++++++++ 1 file changed, 16 insertions(+) diff --git a/src/linear_algebra/span.lean b/src/linear_algebra/span.lean index b3e27235d2fb1..1b35892df4ab6 100644 --- a/src/linear_algebra/span.lean +++ b/src/linear_algebra/span.lean @@ -557,6 +557,22 @@ begin exact ⟨y, ⟨hyd, by simpa only [span_le, singleton_subset_iff]⟩⟩, end +/-- The span of a finite subset is compact in the lattice of submodules. -/ +lemma finset_span_is_compact_element (S : finset M) : + complete_lattice.is_compact_element (span R S : submodule R M) := +begin + rw span_eq_supr_of_singleton_spans, + simp only [finset.mem_coe], + rw ←finset.sup_eq_supr, + exact complete_lattice.finset_sup_compact_of_compact S + (λ x _, singleton_span_is_compact_element x), +end + +/-- The span of a finite subset is compact in the lattice of submodules. -/ +lemma finite_span_is_compact_element (S : set M) (h : S.finite) : + complete_lattice.is_compact_element (span R S : submodule R M) := +finite.coe_to_finset h ▸ (finset_span_is_compact_element h.to_finset) + instance : is_compactly_generated (submodule R M) := ⟨λ s, ⟨(λ x, span R {x}) '' s, ⟨λ t ht, begin rcases (set.mem_image _ _ _).1 ht with ⟨x, hx, rfl⟩, From ceca8d741ffe664b90353d0a82ded9103082ba36 Mon Sep 17 00:00:00 2001 From: negiizhao Date: Wed, 4 May 2022 11:10:47 +0000 Subject: [PATCH 098/153] fix(ring_theory/polynomial/basic): fix unexpected change of an implicit parameter (#13935) Fix unexpected change of an implicit parameter in the previous PR(#13800). Fix docstring. --- src/ring_theory/polynomial/basic.lean | 7 +++---- 1 file changed, 3 insertions(+), 4 deletions(-) diff --git a/src/ring_theory/polynomial/basic.lean b/src/ring_theory/polynomial/basic.lean index aa146e5e2066d..f9abb04851f25 100644 --- a/src/ring_theory/polynomial/basic.lean +++ b/src/ring_theory/polynomial/basic.lean @@ -105,7 +105,7 @@ begin end /-- The first `n` coefficients on `degree_lt n` form a linear equivalence with `fin n → R`. -/ -def degree_lt_equiv (n : ℕ) : degree_lt R n ≃ₗ[R] (fin n → R) := +def degree_lt_equiv (R) [semiring R] (n : ℕ) : degree_lt R n ≃ₗ[R] (fin n → R) := { to_fun := λ p n, (↑p : R[X]).coeff n, inv_fun := λ f, ⟨∑ i : fin n, monomial i (f i), (degree_lt R n).sum_mem (λ i _, mem_degree_lt.mpr (lt_of_le_of_lt @@ -212,8 +212,7 @@ the ring closure of the original coefficients. -/ def restriction (p : R[X]) : polynomial (subring.closure (↑p.frange : set R)) := ∑ i in p.support, monomial i (⟨p.coeff i, if H : p.coeff i = 0 then H.symm ▸ (subring.closure _).zero_mem - else subring.subset_closure (p.coeff_mem_frange _ H)⟩ : - (subring.closure (↑p.frange : set R))) + else subring.subset_closure (p.coeff_mem_frange _ H)⟩ : (subring.closure (↑p.frange : set R))) @[simp] theorem coeff_restriction {p : R[X]} {n : ℕ} : ↑(coeff (restriction p) n) = coeff p n := @@ -277,7 +276,7 @@ section to_subring variables (p : R[X]) (T : subring R) /-- Given a polynomial `p` and a subring `T` that contains the coefficients of `p`, -return the corresponding polynomial whose coefficients are in `T. -/ +return the corresponding polynomial whose coefficients are in `T`. -/ def to_subring (hp : (↑p.frange : set R) ⊆ T) : T[X] := ∑ i in p.support, monomial i (⟨p.coeff i, if H : p.coeff i = 0 then H.symm ▸ T.zero_mem From 32320a100c580e1f5b0339ebf03c2b8d8a92fe90 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Wed, 4 May 2022 14:57:36 +0000 Subject: [PATCH 099/153] feat(measure_theory/integral/lebesgue): speed up a proof (#13946) --- src/measure_theory/integral/lebesgue.lean | 11 +++++++---- 1 file changed, 7 insertions(+), 4 deletions(-) diff --git a/src/measure_theory/integral/lebesgue.lean b/src/measure_theory/integral/lebesgue.lean index 5d76c691b60f5..8f13b3c789654 100644 --- a/src/measure_theory/integral/lebesgue.lean +++ b/src/measure_theory/integral/lebesgue.lean @@ -1329,7 +1329,8 @@ begin begin rw [← simple_func.add_lintegral, ← simple_func.map_add @ennreal.coe_add], refine simple_func.lintegral_mono (λ x, _) le_rfl, - simp [-ennreal.coe_add, add_tsub_eq_max, le_max_right] + simp only [add_tsub_eq_max, le_max_right, coe_map, function.comp_app, simple_func.coe_add, + simple_func.coe_sub, pi.add_apply, pi.sub_apply, with_top.coe_max] end ... ≤ (map coe φ).lintegral (μ.restrict s) + ε₁ : begin @@ -1337,9 +1338,11 @@ begin exact simple_func.lintegral_mono le_rfl measure.restrict_le_self end ... ≤ (simple_func.const α (C : ℝ≥0∞)).lintegral (μ.restrict s) + ε₁ : - by { mono*, exacts [λ x, coe_le_coe.2 (hC x), le_rfl, le_rfl] } - ... = C * μ s + ε₁ : by simp [← simple_func.lintegral_eq_lintegral] - ... ≤ C * ((ε₂ - ε₁) / C) + ε₁ : by { mono*, exacts [le_rfl, hs.le, le_rfl] } + add_le_add (simple_func.lintegral_mono (λ x, coe_le_coe.2 (hC x)) le_rfl) le_rfl + ... = C * μ s + ε₁ : by simp only [←simple_func.lintegral_eq_lintegral, coe_const, + lintegral_const, measure.restrict_apply, measurable_set.univ, univ_inter] + ... ≤ C * ((ε₂ - ε₁) / C) + ε₁ : + add_le_add_right (ennreal.mul_le_mul le_rfl hs.le) _ ... ≤ (ε₂ - ε₁) + ε₁ : add_le_add mul_div_le le_rfl ... = ε₂ : tsub_add_cancel_of_le hε₁₂.le, end From d565adb7ef4acb81b73435eca0aca865fa09a9c5 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 4 May 2022 15:53:23 +0000 Subject: [PATCH 100/153] feat(analysis/convex/topology): Separating by convex sets (#11458) When `s` is compact, `t` is closed and both are convex, we can find disjoint open convex sets containing `s` and `t`. --- src/analysis/convex/basic.lean | 12 +++-- src/analysis/convex/topology.lean | 32 +++++++++++--- .../metric_space/hausdorff_distance.lean | 44 ++++++++++++++++++- 3 files changed, 78 insertions(+), 10 deletions(-) diff --git a/src/analysis/convex/basic.lean b/src/analysis/convex/basic.lean index 162b16ef7cba2..3d5eda6686d7c 100644 --- a/src/analysis/convex/basic.lean +++ b/src/analysis/convex/basic.lean @@ -567,10 +567,14 @@ lemma convex_sInter {S : set (set E)} (h : ∀ s ∈ S, convex 𝕜 s) : convex assume x y hx hy a b ha hb hab s hs, h s hs (hx s hs) (hy s hs) ha hb hab -lemma convex_Inter {ι : Sort*} {s : ι → set E} (h : ∀ i : ι, convex 𝕜 (s i)) : - convex 𝕜 (⋂ i, s i) := +lemma convex_Inter {ι : Sort*} {s : ι → set E} (h : ∀ i, convex 𝕜 (s i)) : convex 𝕜 (⋂ i, s i) := (sInter_range s) ▸ convex_sInter $ forall_range_iff.2 h +lemma convex_Inter₂ {ι : Sort*} {κ : ι → Sort*} {s : Π i, κ i → set E} + (h : ∀ i j, convex 𝕜 (s i j)) : + convex 𝕜 (⋂ i j, s i j) := +convex_Inter $ λ i, convex_Inter $ h i + lemma convex.prod {s : set E} {t : set F} (hs : convex 𝕜 s) (ht : convex 𝕜 t) : convex 𝕜 (s ×ˢ t) := begin @@ -963,8 +967,8 @@ begin rw [convex.combo_affine_apply hab, hx'f, hy'f] end -lemma convex.neg (hs : convex 𝕜 s) : convex 𝕜 ((λ z, -z) '' s) := -hs.is_linear_image is_linear_map.is_linear_map_neg +lemma convex.neg (hs : convex 𝕜 s) : convex 𝕜 (-s) := +by { rw ←set.image_neg, exact hs.is_linear_image is_linear_map.is_linear_map_neg } lemma convex.neg_preimage (hs : convex 𝕜 s) : convex 𝕜 ((λ z, -z) ⁻¹' s) := hs.is_linear_preimage is_linear_map.is_linear_map_neg diff --git a/src/analysis/convex/topology.lean b/src/analysis/convex/topology.lean index 5caddce64ffd5..ab37411b7b782 100644 --- a/src/analysis/convex/topology.lean +++ b/src/analysis/convex/topology.lean @@ -1,9 +1,10 @@ /- Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Alexander Bentkamp, Yury Kudriashov +Authors: Alexander Bentkamp, Yury Kudryashov -/ import analysis.convex.jensen +import analysis.normed.group.pointwise import analysis.normed_space.finite_dimension import analysis.normed_space.ray import topology.path_connected @@ -32,7 +33,7 @@ We prove the following facts: variables {ι : Type*} {E : Type*} -open set +open metric set open_locale pointwise convex lemma real.convex_iff_is_preconnected {s : set ℝ} : convex ℝ s ↔ is_preconnected s := @@ -288,11 +289,11 @@ end has_continuous_smul /-! ### Normed vector space -/ section normed_space -variables [semi_normed_group E] [normed_space ℝ E] +variables [semi_normed_group E] [normed_space ℝ E] {s t : set E} /-- The norm on a real normed space is convex on any convex set. See also `seminorm.convex_on` and `convex_on_univ_norm`. -/ -lemma convex_on_norm {s : set E} (hs : convex ℝ s) : convex_on ℝ s norm := +lemma convex_on_norm (hs : convex ℝ s) : convex_on ℝ s norm := ⟨hs, λ x y hx hy a b ha hb hab, calc ∥a • x + b • y∥ ≤ ∥a • x∥ + ∥b • y∥ : norm_add_le _ _ ... = a * ∥x∥ + b * ∥y∥ @@ -302,7 +303,7 @@ lemma convex_on_norm {s : set E} (hs : convex ℝ s) : convex_on ℝ s norm := and `convex_on_norm`. -/ lemma convex_on_univ_norm : convex_on ℝ univ (norm : E → ℝ) := convex_on_norm convex_univ -lemma convex_on_dist (z : E) {s : set E} (hs : convex ℝ s) : convex_on ℝ s (λz', dist z' z) := +lemma convex_on_dist (z : E) (hs : convex ℝ s) : convex_on ℝ s (λ z', dist z' z) := by simpa [dist_eq_norm, preimage_preimage] using (convex_on_norm (hs.translate (-z))).comp_affine_map (affine_map.id ℝ E - affine_map.const ℝ E z) @@ -316,6 +317,27 @@ by simpa only [metric.ball, sep_univ] using (convex_on_univ_dist a).convex_lt r lemma convex_closed_ball (a : E) (r : ℝ) : convex ℝ (metric.closed_ball a r) := by simpa only [metric.closed_ball, sep_univ] using (convex_on_univ_dist a).convex_le r +lemma convex.thickening (hs : convex ℝ s) (δ : ℝ) : convex ℝ (thickening δ s) := +by { rw ←add_ball_zero, exact hs.add (convex_ball 0 _) } + +lemma convex.cthickening (hs : convex ℝ s) (δ : ℝ) : convex ℝ (cthickening δ s) := +begin + obtain hδ | hδ := le_total 0 δ, + { rw cthickening_eq_Inter_thickening hδ, + exact convex_Inter₂ (λ _ _, hs.thickening _) }, + { rw cthickening_of_nonpos hδ, + exact hs.closure } +end + +/-- If `s`, `t` are disjoint convex sets, `s` is compact and `t` is closed then we can find open +disjoint convex sets containing them. -/ +lemma disjoint.exists_open_convexes (disj : disjoint s t) (hs₁ : convex ℝ s) (hs₂ : is_compact s) + (ht₁ : convex ℝ t) (ht₂ : is_closed t) : + ∃ u v, is_open u ∧ is_open v ∧ convex ℝ u ∧ convex ℝ v ∧ s ⊆ u ∧ t ⊆ v ∧ disjoint u v := +let ⟨δ, hδ, hst⟩ := disj.exists_thickenings hs₂ ht₂ in + ⟨_, _, is_open_thickening, is_open_thickening, hs₁.thickening _, ht₁.thickening _, + self_subset_thickening hδ _, self_subset_thickening hδ _, hst⟩ + /-- Given a point `x` in the convex hull of `s` and a point `y`, there exists a point of `s` at distance at least `dist x y` from `y`. -/ lemma convex_hull_exists_dist_ge {s : set E} {x : E} (hx : x ∈ convex_hull ℝ s) (y : E) : diff --git a/src/topology/metric_space/hausdorff_distance.lean b/src/topology/metric_space/hausdorff_distance.lean index 0dd2c8995a778..f402d7a00abca 100644 --- a/src/topology/metric_space/hausdorff_distance.lean +++ b/src/topology/metric_space/hausdorff_distance.lean @@ -174,6 +174,17 @@ begin exact ⟨y, ys, le_antisymm (inf_edist_le_edist_of_mem ys) (by rwa le_inf_edist)⟩ end +lemma exists_pos_forall_le_edist (hs : is_compact s) (hs' : s.nonempty) (ht : is_closed t) + (hst : disjoint s t) : + ∃ r, 0 < r ∧ ∀ (x ∈ s) (y ∈ t), r ≤ edist x y := +begin + obtain ⟨x, hx, h⟩ : ∃ x ∈ s, ∀ y ∈ s, inf_edist x t ≤ inf_edist y t := + hs.exists_forall_le hs' continuous_inf_edist.continuous_on, + refine ⟨inf_edist x t, pos_iff_ne_zero.2 $ λ H, hst ⟨hx, _⟩, λ y hy, le_inf_edist.1 $ h y hy⟩, + rw ←ht.closure_eq, + exact mem_closure_iff_inf_edist_zero.2 H, +end + end inf_edist --section /-! ### The Hausdorff distance as a function into `ℝ≥0∞`. -/ @@ -877,7 +888,7 @@ end thickening --section section cthickening -variables [pseudo_emetric_space α] {δ ε : ℝ} {s : set α} {x : α} +variables [pseudo_emetric_space α] {δ ε : ℝ} {s t : set α} {x : α} open emetric @@ -1023,6 +1034,37 @@ by simp_rw [thickening, inf_edist_closure] @[simp] lemma cthickening_closure : cthickening δ (closure s) = cthickening δ s := by simp_rw [cthickening, inf_edist_closure] +open ennreal + +lemma _root_.disjoint.exists_thickenings (hst : disjoint s t) (hs : is_compact s) + (ht : is_closed t) : + ∃ δ, 0 < δ ∧ disjoint (thickening δ s) (thickening δ t) := +begin + obtain rfl | hs' := s.eq_empty_or_nonempty, + { simp_rw thickening_empty, + exact ⟨1, zero_lt_one, empty_disjoint _⟩ }, + obtain ⟨r, hr, h⟩ := exists_pos_forall_le_edist hs hs' ht hst, + refine ⟨(min 1 (r/2)).to_real, to_real_pos (lt_min ennreal.zero_lt_one $ half_pos hr.ne').ne' + (min_lt_of_left_lt one_lt_top).ne, _⟩, + rintro z ⟨hzs, hzt⟩, + rw mem_thickening_iff_exists_edist_lt at hzs hzt, + obtain ⟨x, hx, hzx⟩ := hzs, + obtain ⟨y, hy, hzy⟩ := hzt, + refine (((h _ hx _ hy).trans $ edist_triangle_left _ _ _).trans_lt $ + ennreal.add_lt_add hzx hzy).not_le _, + rw ←two_mul, + exact ennreal.mul_le_of_le_div' (of_real_to_real_le.trans $ min_le_right _ _), +end + +lemma _root_.disjoint.exists_cthickenings (hst : disjoint s t) (hs : is_compact s) + (ht : is_closed t) : + ∃ δ, 0 < δ ∧ disjoint (cthickening δ s) (cthickening δ t) := +begin + obtain ⟨δ, hδ, h⟩ := hst.exists_thickenings hs ht, + refine ⟨δ / 2, half_pos hδ, h.mono _ _⟩; + exact (cthickening_subset_thickening' hδ (half_lt_self hδ) _), +end + lemma cthickening_eq_Inter_cthickening' {δ : ℝ} (s : set ℝ) (hsδ : s ⊆ Ioi δ) (hs : ∀ ε, δ < ε → (s ∩ (Ioc δ ε)).nonempty) (E : set α) : cthickening δ E = ⋂ ε ∈ s, cthickening ε E := From a7c50973a983daab8e850d3c97abe8bc4d9ea2e0 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 4 May 2022 15:53:24 +0000 Subject: [PATCH 101/153] feat(set_theory/cardinal/cofinality): Cofinality of normal functions (#12384) If `f` is normal and `a` is limit, then `cof (f a) = cof a`. We use this to golf `cof_add` from 24 lines down to 6. --- src/set_theory/cardinal/cofinality.lean | 83 +++++++++++++++++-------- 1 file changed, 57 insertions(+), 26 deletions(-) diff --git a/src/set_theory/cardinal/cofinality.lean b/src/set_theory/cardinal/cofinality.lean index ebf76d959306b..ba23e02b7568c 100644 --- a/src/set_theory/cardinal/cofinality.lean +++ b/src/set_theory/cardinal/cofinality.lean @@ -345,6 +345,8 @@ bsup_lt_ord_lift (by rwa (o.card).lift_id) (mk_eq_zero_iff.1 (e.trans z)).elim' ⟨_, h⟩⟩, λ e, by simp [e]⟩ +theorem cof_ne_zero {o} : cof o ≠ 0 ↔ o ≠ 0 := cof_eq_zero.not + @[simp] theorem cof_succ (o) : cof (succ o) = 1 := begin apply le_antisymm, @@ -382,32 +384,6 @@ end apply subsingleton.elim } } end, λ ⟨a, e⟩, by simp [e]⟩ -@[simp] theorem cof_add (a b : ordinal) : b ≠ 0 → cof (a + b) = cof b := -induction_on a $ λ α r _, induction_on b $ λ β s _ b0, begin - resetI, - change cof (type _) = _, - refine eq_of_forall_le_iff (λ c, _), - rw [le_cof_type, le_cof_type], - split; intros H S hS, - { refine le_trans (H {a | sum.rec_on a (∅:set α) S} (λ a, _)) ⟨⟨_, _⟩⟩, - { cases a with a b, - { cases type_ne_zero_iff_nonempty.1 b0 with b, - rcases hS b with ⟨b', bs, _⟩, - exact ⟨sum.inr b', bs, by simp⟩ }, - { rcases hS b with ⟨b', bs, h⟩, - exact ⟨sum.inr b', bs, by simp [h]⟩ } }, - { exact λ a, match a with ⟨sum.inr b, h⟩ := ⟨b, h⟩ end }, - { exact λ a b, match a, b with - ⟨sum.inr a, h₁⟩, ⟨sum.inr b, h₂⟩, h := by congr; injection h - end } }, - { refine le_trans (H (sum.inr ⁻¹' S) (λ a, _)) ⟨⟨_, _⟩⟩, - { rcases hS (sum.inr a) with ⟨a'|b', bs, h⟩; simp at h, - { cases h }, { exact ⟨b', bs, h⟩ } }, - { exact λ ⟨a, h⟩, ⟨_, h⟩ }, - { exact λ ⟨a, h₁⟩ ⟨b, h₂⟩ h, - by injection h with h; congr; injection h } } -end - /-- A fundamental sequence for `a` is an increasing sequence of length `o = cof a` that converges at `a`. We provide `o` explicitly in order to avoid type rewrites. -/ def is_fundamental_sequence (a o : ordinal.{u}) (f : Π b < o, ordinal.{u}) : Prop := @@ -509,6 +485,55 @@ begin exact ord_injective ((hf.trans hg).cof_eq.symm) end +protected theorem is_normal.is_fundamental_sequence {f : ordinal.{u} → ordinal.{u}} + (hf : is_normal f) {a o} (ha : is_limit a) {g} (hg : is_fundamental_sequence a o g) : + is_fundamental_sequence (f a) o (λ b hb, f (g b hb)) := +begin + refine ⟨_, λ i j _ _ h, hf.strict_mono (hg.2.1 _ _ h), _⟩, + { rcases exists_lsub_cof (f a) with ⟨ι, f', hf', hι⟩, + rw [←hg.cof_eq, ord_le_ord, ←hι], + suffices : lsub.{u u} (λ i, (Inf {b : ordinal | f' i ≤ f b})) = a, + { rw ←this, + apply cof_lsub_le }, + have H : ∀ i, ∃ b < a, f' i ≤ f b := λ i, begin + have := lt_lsub.{u u} f' i, + rwa [hf', ←is_normal.blsub_eq.{u u} hf ha, lt_blsub_iff] at this + end, + refine le_antisymm (lsub_le (λ i, _)) (le_of_forall_lt (λ b hb, _)), + { rcases H i with ⟨b, hb, hb'⟩, + exact lt_of_le_of_lt (cInf_le' hb') hb }, + { have := hf.strict_mono hb, + rw [←hf', lt_lsub_iff] at this, + cases this with i hi, + rcases H i with ⟨b, _, hb⟩, + exact lt_of_le_of_lt ((le_cInf_iff'' ⟨b, hb⟩).2 + (λ c hc, hf.strict_mono.le_iff_le.1 (hi.trans hc))) (lt_lsub _ i) } }, + { rw @blsub_comp.{u u u} a _ (λ b _, f b) (λ i j hi hj h, hf.strict_mono.monotone h) g hg.2.2, + exact is_normal.blsub_eq.{u u} hf ha } +end + +theorem is_normal.cof_eq {f} (hf : is_normal f) {a} (ha : is_limit a) : cof (f a) = cof a := +let ⟨g, hg⟩ := exists_fundamental_sequence a in + ord_injective (hf.is_fundamental_sequence ha hg).cof_eq + +theorem is_normal.cof_le {f} (hf : is_normal f) (a) : cof a ≤ cof (f a) := +begin + rcases zero_or_succ_or_limit a with rfl | ⟨b, rfl⟩ | ha, + { rw cof_zero, + exact zero_le _ }, + { rw [cof_succ, cardinal.one_le_iff_ne_zero, cof_ne_zero, ←ordinal.pos_iff_ne_zero], + exact (ordinal.zero_le (f b)).trans_lt (hf.1 b) }, + { rw hf.cof_eq ha } +end + +@[simp] theorem cof_add (a b : ordinal) : b ≠ 0 → cof (a + b) = cof b := +λ h, begin + rcases zero_or_succ_or_limit b with rfl | ⟨c, rfl⟩ | hb, + { contradiction }, + { rw [add_succ, cof_succ, cof_succ] }, + { exact (add_is_normal a).cof_eq hb } +end + theorem omega_le_cof {o} : ω ≤ cof o ↔ is_limit o := begin rcases zero_or_succ_or_limit o with rfl|⟨o,rfl⟩|l, @@ -526,6 +551,12 @@ begin exact not_succ_is_limit _ l } } end +@[simp] theorem aleph'_cof {o : ordinal} (ho : o.is_limit) : (aleph' o).ord.cof = o.cof := +aleph'_is_normal.cof_eq ho + +@[simp] theorem aleph_cof {o : ordinal} (ho : o.is_limit) : (aleph o).ord.cof = o.cof := +aleph_is_normal.cof_eq ho + @[simp] theorem cof_omega : cof omega = ω := le_antisymm (by rw ← card_omega; apply cof_le_card) From e71613938c5bf2466f931e8c935505ac3929ca64 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 4 May 2022 15:53:26 +0000 Subject: [PATCH 102/153] feat(algebra/homology/Module): API for complexes of modules (#12622) API for homological complexes in `Module R`. Co-authored-by: Scott Morrison Co-authored-by: Kevin Buzzard --- src/algebra/category/Module/kernels.lean | 5 + src/algebra/category/Module/subobject.lean | 30 ++++++ src/algebra/homology/Module.lean | 93 +++++++++++++++++++ src/algebra/homology/homology.lean | 14 +-- .../limits/concrete_category.lean | 10 ++ 5 files changed, 141 insertions(+), 11 deletions(-) create mode 100644 src/algebra/homology/Module.lean diff --git a/src/algebra/category/Module/kernels.lean b/src/algebra/category/Module/kernels.lean index 6ec79a873ac7f..558c9668b0878 100644 --- a/src/algebra/category/Module/kernels.lean +++ b/src/algebra/category/Module/kernels.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Markus Himmel -/ import algebra.category.Module.epi_mono +import category_theory.limits.concrete_category /-! # The concrete (co)kernels in the category of modules are (co)kernels in the categorical sense. @@ -102,4 +103,8 @@ by { convert colimit.iso_colimit_cocone_ι_hom _ _; refl, } ↿f.range.mkq ≫ (cokernel_iso_range_quotient f).inv = cokernel.π f := by { convert colimit.iso_colimit_cocone_ι_inv ⟨_, cokernel_is_colimit f⟩ _; refl, } +lemma cokernel_π_ext {M N : Module.{u} R} (f : M ⟶ N) {x y : N} (m : M) (w : x = y + f m) : + cokernel.π f x = cokernel.π f y := +by { subst w, simp, } + end Module diff --git a/src/algebra/category/Module/subobject.lean b/src/algebra/category/Module/subobject.lean index 6c944c9d57804..226ba8b43b876 100644 --- a/src/algebra/category/Module/subobject.lean +++ b/src/algebra/category/Module/subobject.lean @@ -4,7 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Markus Himmel -/ import algebra.category.Module.epi_mono +import algebra.category.Module.kernels import category_theory.subobject.well_powered +import category_theory.subobject.limits +import category_theory.limits.concrete_category /-! # Subobjects in the category of `R`-modules @@ -16,6 +19,8 @@ and its submodules. This immediately implies that the category of `R`-modules is open category_theory open category_theory.subobject +open category_theory.limits + open_locale Module universes v u @@ -63,4 +68,29 @@ noncomputable def subobject_Module : subobject M ≃o submodule R M := order_iso instance well_powered_Module : well_powered (Module.{v} R) := ⟨λ M, ⟨⟨_, ⟨(subobject_Module M).to_equiv⟩⟩⟩⟩ +local attribute [instance] has_kernels_Module + +/-- Bundle an element `m : M` such that `f m = 0` as a term of `kernel_subobject f`. -/ +noncomputable def to_kernel_subobject {M N : Module R} {f : M ⟶ N} : + linear_map.ker f →ₗ[R] kernel_subobject f := +(kernel_subobject_iso f ≪≫ Module.kernel_iso_ker f).inv + +@[simp] lemma to_kernel_subobject_arrow {M N : Module R} {f : M ⟶ N} (x : linear_map.ker f) : + (kernel_subobject f).arrow (to_kernel_subobject x) = x.1 := +by simp [to_kernel_subobject] + +/-- +An extensionality lemma showing that two elements of a cokernel by an image +are equal if they differ by an element of the image. + +The application is for homology: +two elements in homology are equal if they differ by a boundary. +-/ +@[ext] +lemma cokernel_π_image_subobject_ext {L M N : Module.{v} R} + (f : L ⟶ M) [has_image f] (g : (image_subobject f : Module.{v} R) ⟶ N) [has_cokernel g] + {x y : N} (l : L) (w : x = y + g (factor_thru_image_subobject f l)) : + cokernel.π g x = cokernel.π g y := +by { subst w, simp, } + end Module diff --git a/src/algebra/homology/Module.lean b/src/algebra/homology/Module.lean new file mode 100644 index 0000000000000..6f774f21b7f92 --- /dev/null +++ b/src/algebra/homology/Module.lean @@ -0,0 +1,93 @@ +/- +Copyright (c) 2021 Scott Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Scott Morrison +-/ +import algebra.homology.homotopy +import algebra.category.Module.abelian +import algebra.category.Module.subobject + +/-! +# Complexes of modules + +We provide some additional API to work with homological complexes in `Module R`. +-/ + +universes v u + +open_locale classical +noncomputable theory + +open category_theory category_theory.limits homological_complex + +variables {R : Type v} [ring R] +variables {ι : Type*} {c : complex_shape ι} {C D : homological_complex (Module.{u} R) c} + +namespace Module + +/-- +To prove that two maps out of a homology group are equal, +it suffices to check they are equal on the images of cycles. +-/ +lemma homology_ext {L M N K : Module R} {f : L ⟶ M} {g : M ⟶ N} (w : f ≫ g = 0) + {h k : homology f g w ⟶ K} + (w : ∀ (x : linear_map.ker g), + h (cokernel.π (image_to_kernel _ _ w) (to_kernel_subobject x)) = + k (cokernel.π (image_to_kernel _ _ w) (to_kernel_subobject x))) : h = k := +begin + refine cokernel_funext (λ n, _), + -- Gosh it would be nice if `equiv_rw` could directly use an isomorphism, or an enriched `≃`. + equiv_rw (kernel_subobject_iso g ≪≫ Module.kernel_iso_ker g).to_linear_equiv.to_equiv at n, + convert w n; simp [to_kernel_subobject], +end + +/-- Bundle an element `C.X i` such that `C.d_from i x = 0` as a term of `C.cycles i`. -/ +abbreviation to_cycles {C : homological_complex (Module.{u} R) c} + {i : ι} (x : linear_map.ker (C.d_from i)) : C.cycles i := +to_kernel_subobject x + +@[ext] +lemma cycles_ext {C : homological_complex (Module.{u} R) c} {i : ι} + {x y : C.cycles i} (w : (C.cycles i).arrow x = (C.cycles i).arrow y) : x = y := +begin + apply_fun (C.cycles i).arrow using (Module.mono_iff_injective _).mp (cycles C i).arrow_mono, + exact w, +end + +local attribute [instance] concrete_category.has_coe_to_sort + +@[simp] lemma cycles_map_to_cycles (f : C ⟶ D) {i : ι} (x : linear_map.ker (C.d_from i)) : + (cycles_map f i) (to_cycles x) = to_cycles ⟨f.f i x.1, by simp [x.2]⟩ := +by { ext, simp, } + +/-- Build a term of `C.homology i` from an element `C.X i` such that `C.d_from i x = 0`. -/ +abbreviation to_homology + {C : homological_complex (Module.{u} R) c} {i : ι} (x : linear_map.ker (C.d_from i)) : + C.homology i := +homology.π (C.d_to i) (C.d_from i) _ (to_cycles x) + +@[ext] +lemma homology_ext' {M : Module R} (i : ι) {h k : C.homology i ⟶ M} + (w : ∀ (x : linear_map.ker (C.d_from i)), h (to_homology x) = k (to_homology x)) : + h = k := +homology_ext _ w + +/-- We give an alternative proof of `homology_map_eq_of_homotopy`, +specialized to the setting of `V = Module R`, +to demonstrate the use of extensionality lemmas for homology in `Module R`. -/ +example (f g : C ⟶ D) (h : homotopy f g) (i : ι) : + (homology_functor (Module.{u} R) c i).map f = (homology_functor (Module.{u} R) c i).map g := +begin + -- To check that two morphisms out of a homology group agree, it suffices to check on cycles: + ext, + dsimp, + simp only [homology.π_map_apply], + -- To check that two elements are equal mod boundaries, it suffices to exhibit a boundary: + ext1, + swap, exact (to_prev i h.hom) x.1, + -- Moreover, to check that two cycles are equal, it suffices to check their underlying elements: + ext1, + simp [h.comm i, x.2]; abel, +end + +end Module diff --git a/src/algebra/homology/homology.lean b/src/algebra/homology/homology.lean index 2a84e5ad47ca1..786a74f5d3a9c 100644 --- a/src/algebra/homology/homology.lean +++ b/src/algebra/homology/homology.lean @@ -39,13 +39,9 @@ section cycles variables [has_kernels V] /-- The cycles at index `i`, as a subobject. -/ -def cycles (i : ι) : subobject (C.X i) := +abbreviation cycles (i : ι) : subobject (C.X i) := kernel_subobject (C.d_from i) -@[simp, reassoc] -lemma cycles_arrow_d_from (i : ι) : (C.cycles i).arrow ≫ C.d_from i = 0 := -by { dsimp [cycles], simp, } - lemma cycles_eq_kernel_subobject {i j : ι} (r : c.rel i j) : C.cycles i = kernel_subobject (C.d i j) := C.kernel_from_eq_kernel r @@ -116,11 +112,6 @@ image_to_kernel _ _ (C.d_to_comp_d_from i) (C.boundaries i).of_le (C.cycles i) h = C.boundaries_to_cycles i := rfl -@[simp, reassoc] -lemma boundaries_to_cycles_arrow (C : homological_complex V c) (i : ι) : - C.boundaries_to_cycles i ≫ (C.cycles i).arrow = (C.boundaries i).arrow := -by { dsimp [cycles], simp, } - variables [has_cokernels V] /-- @@ -146,7 +137,8 @@ The morphism between cycles induced by a chain map. abbreviation cycles_map (f : C₁ ⟶ C₂) (i : ι) : (C₁.cycles i : V) ⟶ (C₂.cycles i : V) := subobject.factor_thru _ ((C₁.cycles i).arrow ≫ f.f i) (kernel_subobject_factors _ _ (by simp)) -@[simp] lemma cycles_map_arrow (f : C₁ ⟶ C₂) (i : ι) : +@[simp, reassoc, elementwise] +lemma cycles_map_arrow (f : C₁ ⟶ C₂) (i : ι) : (cycles_map f i) ≫ (C₂.cycles i).arrow = (C₁.cycles i).arrow ≫ f.f i := by { simp, } diff --git a/src/category_theory/limits/concrete_category.lean b/src/category_theory/limits/concrete_category.lean index 143419763afda..6c136a2d60c0f 100644 --- a/src/category_theory/limits/concrete_category.lean +++ b/src/category_theory/limits/concrete_category.lean @@ -149,6 +149,16 @@ end limits section colimits +-- We don't mark this as an `@[ext]` lemma as we don't always want to work elementwise. +lemma cokernel_funext {C : Type*} [category C] [has_zero_morphisms C] [concrete_category C] + {M N K : C} {f : M ⟶ N} [has_cokernel f] {g h : cokernel f ⟶ K} + (w : ∀ (n : N), g (cokernel.π f n) = h (cokernel.π f n)) : g = h := +begin + apply coequalizer.hom_ext, + apply concrete_category.hom_ext _ _, + simpa using w, +end + variables {C : Type u} [category.{v} C] [concrete_category.{v} C] {J : Type v} [small_category J] (F : J ⥤ C) [preserves_colimit F (forget C)] From 552a470943af7f335184c272deac01eac0237df9 Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Wed, 4 May 2022 15:53:27 +0000 Subject: [PATCH 103/153] feat(number_theory/cyclotomic/rat): the ring of integers of cyclotomic fields. (#13585) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We compute the ring of integers of a `p ^ n`-th cyclotomic extension of `ℚ`. From flt-regular Co-authored-by: Eric Rodriguez --- src/number_theory/cyclotomic/basic.lean | 5 + .../cyclotomic/discriminant.lean | 85 ++++++++- src/number_theory/cyclotomic/rat.lean | 166 ++++++++++++++++++ .../polynomial/cyclotomic/basic.lean | 14 ++ 4 files changed, 265 insertions(+), 5 deletions(-) create mode 100644 src/number_theory/cyclotomic/rat.lean diff --git a/src/number_theory/cyclotomic/basic.lean b/src/number_theory/cyclotomic/basic.lean index a6aadf1123f49..df2ab5d3a0fb8 100644 --- a/src/number_theory/cyclotomic/basic.lean +++ b/src/number_theory/cyclotomic/basic.lean @@ -99,6 +99,11 @@ by simp [is_cyclotomic_extension_iff] lemma empty [h : is_cyclotomic_extension ∅ A B] : (⊥ : subalgebra A B) = ⊤ := by simpa [algebra.eq_top_iff, is_cyclotomic_extension_iff] using h +/-- If `is_cyclotomic_extension {1} A B`, then `A = B`. -/ +lemma singleton_one [h : is_cyclotomic_extension {1} A B] : (⊥ : subalgebra A B) = ⊤ := +algebra.eq_top_iff.2 (λ x, by simpa [adjoin_singleton_one] + using ((is_cyclotomic_extension_iff _ _ _).1 h).2 x) + /-- Transitivity of cyclotomic extensions. -/ lemma trans (C : Type w) [comm_ring C] [algebra A C] [algebra B C] [is_scalar_tower A B C] [hS : is_cyclotomic_extension S A B] diff --git a/src/number_theory/cyclotomic/discriminant.lean b/src/number_theory/cyclotomic/discriminant.lean index 1c8bb0fe92095..6632b37d9aaa2 100644 --- a/src/number_theory/cyclotomic/discriminant.lean +++ b/src/number_theory/cyclotomic/discriminant.lean @@ -9,7 +9,7 @@ import ring_theory.discriminant /-! # Discriminant of cyclotomic fields -We compute the discriminant of a `p`-th cyclotomic extension. +We compute the discriminant of a `p ^ n`-th cyclotomic extension. ## Main results * `is_cyclotomic_extension.discr_odd_prime` : if `p` is an odd prime such that @@ -58,15 +58,15 @@ variables [algebra K L] /-- If `p` is a prime and `is_cyclotomic_extension {p ^ (k + 1)} K L`, then the discriminant of `hζ.power_basis K` is `(-1) ^ ((p ^ (k + 1).totient) / 2) * p ^ (p ^ k * ((p - 1) * (k + 1) - 1))` -if `irreducible (cyclotomic (p ^ (k + 1)) K))`, `irreducible (cyclotomic p K)` and -`p ^ (k + 1) ≠ 2`. -/ +if `irreducible (cyclotomic (p ^ (k + 1)) K))`, and `p ^ (k + 1) ≠ 2`. -/ lemma discr_prime_pow_ne_two [is_cyclotomic_extension {p ^ (k + 1)} K L] [hp : fact (p : ℕ).prime] [ne_zero ((p : ℕ) : K)] (hζ : is_primitive_root ζ ↑(p ^ (k + 1))) (hirr : irreducible (cyclotomic (↑(p ^ (k + 1)) : ℕ) K)) - (hirr₁ : irreducible (cyclotomic (p : ℕ) K)) (hk : p ^ (k + 1) ≠ 2) : + (hk : p ^ (k + 1) ≠ 2) : discr K (hζ.power_basis K).basis = (-1) ^ (((p ^ (k + 1) : ℕ).totient) / 2) * p ^ ((p : ℕ) ^ k * ((p - 1) * (k + 1) - 1)) := begin + have hirr₁ := cyclotomic_irreducible_of_irreducible_pow hp.1 k.succ_ne_zero hirr, haveI : ne_zero ((↑(p ^ (k + 1)) : ℕ) : K), { refine ⟨λ hzero, _⟩, rw [pnat.pow_coe] at hzero, @@ -134,6 +134,81 @@ begin { apply_instance } end +/-- If `p` is a prime and `is_cyclotomic_extension {p ^ (k + 1)} K L`, then the discriminant of +`hζ.power_basis K` is `(-1) ^ (p ^ k * (p - 1) / 2) * p ^ (p ^ k * ((p - 1) * (k + 1) - 1))` +if `irreducible (cyclotomic (p ^ (k + 1)) K))`, and `p ^ (k + 1) ≠ 2`. -/ +lemma discr_prime_pow_ne_two' [is_cyclotomic_extension {p ^ (k + 1)} K L] [hp : fact (p : ℕ).prime] + [ne_zero ((p : ℕ) : K)] (hζ : is_primitive_root ζ ↑(p ^ (k + 1))) + (hirr : irreducible (cyclotomic (↑(p ^ (k + 1)) : ℕ) K)) (hk : p ^ (k + 1) ≠ 2) : + discr K (hζ.power_basis K).basis = + (-1) ^ (((p : ℕ) ^ k * (p - 1)) / 2) * p ^ ((p : ℕ) ^ k * ((p - 1) * (k + 1) - 1)) := +by simpa [totient_prime_pow hp.out (succ_pos k)] using discr_prime_pow_ne_two hζ hirr hk + +/-- If `p` is a prime and `is_cyclotomic_extension {p ^ k} K L`, then the discriminant of +`hζ.power_basis K` is `(-1) ^ ((p ^ k).totient / 2) * p ^ (p ^ (k - 1) * ((p - 1) * k - 1))` +if `irreducible (cyclotomic (p ^ k) K))`. Beware that in the cases `p ^ k = 1` and `p ^ k = 2` +the formula uses `1 / 2 = 0` and `0 - 1 = 0`. It is useful only to have a uniform result. +See also `is_cyclotomic_extension.discr_prime_pow_eq_unit_mul_pow`. -/ +lemma discr_prime_pow [hcycl : is_cyclotomic_extension {p ^ k} K L] [hp : fact (p : ℕ).prime] + [ne_zero ((p : ℕ) : K)] (hζ : is_primitive_root ζ ↑(p ^ k)) + (hirr : irreducible (cyclotomic (↑(p ^ k) : ℕ) K)) : + discr K (hζ.power_basis K).basis = + (-1) ^ (((p ^ k : ℕ).totient) / 2) * p ^ ((p : ℕ) ^ (k - 1) * ((p - 1) * k - 1)) := +begin + unfreezingI { cases k }, + { haveI : ne_zero ((↑(p ^ 0) : ℕ) : K) := ⟨by simp⟩, + simp only [coe_basis, pow_zero, power_basis_gen, totient_one, mul_zero, mul_one, show 1 / 2 = 0, + by refl, discr, trace_matrix], + have hζone : ζ = 1 := by simpa using hζ, + rw [hζ.power_basis_dim _, hζone, ← (algebra_map K L).map_one, + minpoly.eq_X_sub_C_of_algebra_map_inj _ (algebra_map K L).injective, nat_degree_X_sub_C], + simp only [trace_matrix, map_one, one_pow, matrix.det_unique, trace_form_apply, mul_one], + rw [← (algebra_map K L).map_one, trace_algebra_map, finrank _ hirr], + { simp }, + { simpa using hcycl } }, + { by_cases hk : p ^ (k + 1) = 2, + { haveI : ne_zero ((↑(p ^ (k + 1)) : ℕ) : K), + { refine ⟨λ hzero, _⟩, + rw [pnat.pow_coe] at hzero, + simpa [ne_zero.ne ((p : ℕ) : K)] using hzero }, + have hp : p = 2, + { rw [← pnat.coe_inj, pnat.coe_bit0, pnat.one_coe, pnat.pow_coe, ← pow_one 2] at hk, + replace hk := eq_of_prime_pow_eq (prime_iff.1 hp.out) (prime_iff.1 nat.prime_two) + (succ_pos _) hk, + rwa [show 2 = ((2 : ℕ+) : ℕ), by simp, pnat.coe_inj] at hk }, + rw [hp, ← pnat.coe_inj, pnat.pow_coe, pnat.coe_bit0, pnat.one_coe] at hk, + nth_rewrite 1 [← pow_one 2] at hk, + replace hk := nat.pow_right_injective rfl.le hk, + rw [add_left_eq_self] at hk, + simp only [hp, hk, pow_one, pnat.coe_bit0, pnat.one_coe] at hζ, + simp only [hp, hk, show 1 / 2 = 0, by refl, coe_basis, pow_one, power_basis_gen, + pnat.coe_bit0, pnat.one_coe, totient_two, pow_zero, mul_one, mul_zero], + rw [power_basis_dim, hζ.eq_neg_one_of_two_right, show (-1 : L) = algebra_map K L (-1), + by simp, minpoly.eq_X_sub_C_of_algebra_map_inj _ (algebra_map K L).injective, + nat_degree_X_sub_C], + simp only [discr, trace_matrix, matrix.det_unique, fin.default_eq_zero, fin.coe_zero, + pow_zero, trace_form_apply, mul_one], + rw [← (algebra_map K L).map_one, trace_algebra_map, finrank _ hirr, hp, hk], + { simp }, + { apply_instance } }, + { exact discr_prime_pow_ne_two hζ hirr hk } } +end + +/-- If `p` is a prime and `is_cyclotomic_extension {p ^ k} K L`, then there are `u : ℤˣ` and +`n : ℕ` such that the discriminant of `hζ.power_basis K` is `u * p ^ n`. Often this is enough and +less cumbersome to use than `is_cyclotomic_extension.discr_prime_pow`. -/ +lemma discr_prime_pow_eq_unit_mul_pow [is_cyclotomic_extension {p ^ k} K L] + [hp : fact (p : ℕ).prime] [ne_zero ((p : ℕ) : K)] (hζ : is_primitive_root ζ ↑(p ^ k)) + (hirr : irreducible (cyclotomic (↑(p ^ k) : ℕ) K)) : + ∃ (u : ℤˣ) (n : ℕ), discr K (hζ.power_basis K).basis = u * p ^ n := +begin + rw [discr_prime_pow hζ hirr], + by_cases heven : even (((p ^ k : ℕ).totient) / 2), + { refine ⟨1, (p : ℕ) ^ (k - 1) * ((p - 1) * k - 1), by simp [heven.neg_one_pow]⟩ }, + { exact ⟨-1, (p : ℕ) ^ (k - 1) * ((p - 1) * k - 1), + by simp [(odd_iff_not_even.2 heven).neg_one_pow]⟩ }, +end + /-- If `p` is an odd prime and `is_cyclotomic_extension {p} K L`, then `discr K (hζ.power_basis K).basis = (-1) ^ ((p - 1) / 2) * p ^ (p - 2)` if `irreducible (cyclotomic p K)`. -/ @@ -146,7 +221,7 @@ begin { rw [zero_add, pow_one], apply_instance }, have hζ' : is_primitive_root ζ ↑(p ^ (0 + 1)) := by simpa using hζ, - convert discr_prime_pow_ne_two hζ' (by simpa [hirr]) (by simp [hirr]) (by simp [hodd]), + convert discr_prime_pow_ne_two hζ' (by simpa [hirr]) (by simp [hodd]), { rw [zero_add, pow_one, totient_prime hp.out] }, { rw [pow_zero, one_mul, zero_add, mul_one, nat.sub_sub] } end diff --git a/src/number_theory/cyclotomic/rat.lean b/src/number_theory/cyclotomic/rat.lean new file mode 100644 index 0000000000000..d02a490ee2967 --- /dev/null +++ b/src/number_theory/cyclotomic/rat.lean @@ -0,0 +1,166 @@ +/- +Copyright (c) 2022 Riccardo Brasca. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Riccardo Brasca +-/ + +import number_theory.cyclotomic.discriminant +import ring_theory.polynomial.eisenstein + +/-! +# Ring of integers of `p ^ n`-th cyclotomic fields +We compute the ring of integers of a `p ^ n`-th cyclotomic extension of `ℚ`. + +## Main results +* `is_cyclotomic_extension.rat.is_integral_closure_adjoing_singleton_of_prime_pow`: if `K` is a + `p ^ k`-th cyclotomic extension of `ℚ`, then `(adjoin ℤ {ζ})` is the integral closure of + `ℤ` in `K`. +* `is_cyclotomic_extension.rat.cyclotomic_ring_is_integral_closure_of_prime_pow`: the integral + closure of `ℤ` inside `cyclotomic_field (p ^ k) ℚ` is `cyclotomic_ring (p ^ k) ℤ ℚ`. +-/ + +universes u + +open algebra is_cyclotomic_extension polynomial + +open_locale cyclotomic + +namespace is_cyclotomic_extension.rat + +variables {p : ℕ+} {k : ℕ} {K : Type u} [field K] [char_zero K] {ζ : K} [hp : fact (p : ℕ).prime] + +include hp + +/-- The discriminant of the power basis given by `ζ - 1`. -/ +lemma discr_prime_pow_ne_two' [is_cyclotomic_extension {p ^ (k + 1)} ℚ K] + (hζ : is_primitive_root ζ ↑(p ^ (k + 1))) (hk : p ^ (k + 1) ≠ 2) : + discr ℚ (hζ.sub_one_power_basis ℚ).basis = + (-1) ^ (((p ^ (k + 1) : ℕ).totient) / 2) * p ^ ((p : ℕ) ^ k * ((p - 1) * (k + 1) - 1)) := +begin + rw [← discr_prime_pow_ne_two hζ (cyclotomic.irreducible_rat (p ^ (k + 1)).pos) hk], + exact hζ.discr_zeta_eq_discr_zeta_sub_one.symm +end + +lemma discr_odd_prime' [is_cyclotomic_extension {p} ℚ K] (hζ : is_primitive_root ζ p) + (hodd : p ≠ 2) : + discr ℚ (hζ.sub_one_power_basis ℚ).basis = (-1) ^ (((p : ℕ) - 1) / 2) * p ^ ((p : ℕ) - 2) := +begin + rw [← discr_odd_prime hζ (cyclotomic.irreducible_rat hp.out.pos) hodd], + exact hζ.discr_zeta_eq_discr_zeta_sub_one.symm +end + +/-- The discriminant of the power basis given by `ζ - 1`. Beware that in the cases `p ^ k = 1` and +`p ^ k = 2` the formula uses `1 / 2 = 0` and `0 - 1 = 0`. It is useful only to have a uniform +result. See also `is_cyclotomic_extension.rat.discr_prime_pow_eq_unit_mul_pow'`. -/ +lemma discr_prime_pow' [is_cyclotomic_extension {p ^ k} ℚ K] + (hζ : is_primitive_root ζ ↑(p ^ k)) : + discr ℚ (hζ.sub_one_power_basis ℚ).basis = + (-1) ^ (((p ^ k : ℕ).totient) / 2) * p ^ ((p : ℕ) ^ (k - 1) * ((p - 1) * k - 1)) := +begin + rw [← discr_prime_pow hζ (cyclotomic.irreducible_rat (p ^ k).pos)], + exact hζ.discr_zeta_eq_discr_zeta_sub_one.symm +end + +/-- If `p` is a prime and `is_cyclotomic_extension {p ^ k} K L`, then there are `u : ℤˣ` and +`n : ℕ` such that the discriminant of the power basis given by `ζ - 1` is `u * p ^ n`. Often this is +enough and less cumbersome to use than `is_cyclotomic_extension.rat.discr_prime_pow'`. -/ +lemma discr_prime_pow_eq_unit_mul_pow' [is_cyclotomic_extension {p ^ k} ℚ K] + (hζ : is_primitive_root ζ ↑(p ^ k)) : + ∃ (u : ℤˣ) (n : ℕ), discr ℚ (hζ.sub_one_power_basis ℚ).basis = u * p ^ n := +begin + rw [hζ.discr_zeta_eq_discr_zeta_sub_one.symm], + exact discr_prime_pow_eq_unit_mul_pow hζ (cyclotomic.irreducible_rat (p ^ k).pos) +end + +/-- If `K` is a `p ^ k`-th cyclotomic extension of `ℚ`, then `(adjoin ℤ {ζ})` is the +integral closure of `ℤ` in `K`. -/ +lemma is_integral_closure_adjoing_singleton_of_prime_pow + [hcycl : is_cyclotomic_extension {p ^ k} ℚ K] (hζ : is_primitive_root ζ ↑(p ^ k)) : + is_integral_closure (adjoin ℤ ({ζ} : set K)) ℤ K := +begin + refine ⟨subtype.val_injective, λ x, ⟨λ h, ⟨⟨x, _⟩, rfl⟩, _⟩⟩, + swap, + { rintro ⟨y, rfl⟩, + exact is_integral.algebra_map (le_integral_closure_iff_is_integral.1 + (adjoin_le_integral_closure (hζ.is_integral (p ^ k).pos)) _) }, + let B := hζ.sub_one_power_basis ℚ, + have hint : is_integral ℤ B.gen := is_integral_sub (hζ.is_integral (p ^ k).pos) + is_integral_one, + have H := discr_mul_is_integral_mem_adjoin ℚ hint h, + obtain ⟨u, n, hun⟩ := discr_prime_pow_eq_unit_mul_pow' hζ, + rw [hun] at H, + replace H := subalgebra.smul_mem _ H u.inv, + rw [← smul_assoc, ← smul_mul_assoc, units.inv_eq_coe_inv, coe_coe, zsmul_eq_mul, + ← int.cast_mul, units.inv_mul, int.cast_one, one_mul, + show (p : ℚ) ^ n • x = ((p : ℕ) : ℤ) ^ n • x, by simp [smul_def]] at H, + unfreezingI { cases k }, + { haveI : is_cyclotomic_extension {1} ℚ K := by simpa using hcycl, + have : x ∈ (⊥ : subalgebra ℚ K), + { rw [singleton_one ℚ K], + exact mem_top }, + obtain ⟨y, rfl⟩ := mem_bot.1 this, + replace h := (is_integral_algebra_map_iff (algebra_map ℚ K).injective).1 h, + obtain ⟨z, hz⟩ := is_integrally_closed.is_integral_iff.1 h, + rw [← hz, ← is_scalar_tower.algebra_map_apply], + exact subalgebra.algebra_map_mem _ _ }, + { have hmin : (minpoly ℤ B.gen).is_eisenstein_at (submodule.span ℤ {((p : ℕ) : ℤ)}), + { have h₁ := minpoly.gcd_domain_eq_field_fractions ℚ hint, + have h₂ := hζ.minpoly_sub_one_eq_cyclotomic_comp + (cyclotomic.irreducible_rat (p ^ _).pos), + rw [is_primitive_root.sub_one_power_basis_gen] at h₁, + rw [h₁, ← map_cyclotomic_int, show int.cast_ring_hom ℚ = algebra_map ℤ ℚ, by refl, + show ((X + 1)) = map (algebra_map ℤ ℚ) (X + 1), by simp, ← map_comp] at h₂, + rw [is_primitive_root.sub_one_power_basis_gen, map_injective (algebra_map ℤ ℚ) + ((algebra_map ℤ ℚ).injective_int) h₂], + exact cyclotomic_prime_pow_comp_X_add_one_is_eisenstein_at _ _ }, + refine adjoin_le _ (mem_adjoin_of_smul_prime_pow_smul_of_minpoly_is_eiseinstein_at + (nat.prime_iff_prime_int.1 hp.out) hint h H hmin), + simp only [set.singleton_subset_iff, set_like.mem_coe], + exact subalgebra.sub_mem _ (self_mem_adjoin_singleton ℤ _) (subalgebra.one_mem _) } +end + +lemma is_integral_closure_adjoing_singleton_of_prime [hcycl : is_cyclotomic_extension {p} ℚ K] + (hζ : is_primitive_root ζ ↑p) : + is_integral_closure (adjoin ℤ ({ζ} : set K)) ℤ K := +begin + rw [← pow_one p] at hζ hcycl, + exactI is_integral_closure_adjoing_singleton_of_prime_pow hζ, +end + +local attribute [-instance] cyclotomic_field.algebra +local attribute [instance] algebra_rat_subsingleton + +/-- The integral closure of `ℤ` inside `cyclotomic_field (p ^ k) ℚ` is +`cyclotomic_ring (p ^ k) ℤ ℚ`. -/ +lemma cyclotomic_ring_is_integral_closure_of_prime_pow : + is_integral_closure (cyclotomic_ring (p ^ k) ℤ ℚ) ℤ (cyclotomic_field (p ^ k) ℚ) := +begin + haveI : is_cyclotomic_extension {p ^ k} ℚ (cyclotomic_field (p ^ k) ℚ), + { convert cyclotomic_field.is_cyclotomic_extension (p ^ k) _, + { exact subsingleton.elim _ _ }, + { exact ne_zero.char_zero } }, + have hζ := zeta_primitive_root (p ^ k) ℚ (cyclotomic_field (p ^ k) ℚ), + refine ⟨is_fraction_ring.injective _ _, λ x, ⟨λ h, ⟨⟨x, _⟩, rfl⟩, _⟩⟩, + { have := (is_integral_closure_adjoing_singleton_of_prime_pow hζ).is_integral_iff, + obtain ⟨y, rfl⟩ := this.1 h, + convert adjoin_mono _ y.2, + { simp only [eq_iff_true_of_subsingleton] }, + { simp only [eq_iff_true_of_subsingleton] }, + { simp only [pnat.pow_coe, set.singleton_subset_iff, set.mem_set_of_eq], + exact hζ.pow_eq_one } }, + { haveI : is_cyclotomic_extension {p ^ k} ℤ (cyclotomic_ring (p ^ k) ℤ ℚ), + { convert cyclotomic_ring.is_cyclotomic_extension _ ℤ ℚ, + { exact subsingleton.elim _ _ }, + { exact ne_zero.char_zero } }, + rintro ⟨y, rfl⟩, + exact is_integral.algebra_map ((is_cyclotomic_extension.integral {p ^ k} ℤ _) _) } +end + +lemma cyclotomic_ring_is_integral_closure_of_prime : + is_integral_closure (cyclotomic_ring p ℤ ℚ) ℤ (cyclotomic_field p ℚ) := +begin + rw [← pow_one p], + exact cyclotomic_ring_is_integral_closure_of_prime_pow +end + +end is_cyclotomic_extension.rat diff --git a/src/ring_theory/polynomial/cyclotomic/basic.lean b/src/ring_theory/polynomial/cyclotomic/basic.lean index 44945f4109db2..c5bf4d7da35c5 100644 --- a/src/ring_theory/polynomial/cyclotomic/basic.lean +++ b/src/ring_theory/polynomial/cyclotomic/basic.lean @@ -956,6 +956,20 @@ begin nat.totient_mul_of_prime_of_dvd hp hdiv, mul_comm] } end +/-- If `irreducible (cyclotomic (p ^ n) R)` then `irreducible (cyclotomic p R).` -/ +--TODO: add `irreducible (cyclotomic (p ^ m) R)` if `m ≤ n`. +lemma cyclotomic_irreducible_of_irreducible_pow {p : ℕ} (hp : nat.prime p) + {R} [comm_ring R] [is_domain R] : + ∀ {n : ℕ}, n ≠ 0 → irreducible (cyclotomic (p ^ n) R) → irreducible (cyclotomic p R) +| 0 hn _ := by contradiction +| 1 hn h := by rwa pow_one at h +| (n+2) hn h := +begin + rw [pow_succ', ←cyclotomic_expand_eq_cyclotomic hp $ dvd_pow_self p n.succ_ne_zero] at h, + replace h := of_irreducible_expand hp.ne_zero h, + refine cyclotomic_irreducible_of_irreducible_pow n.succ_ne_zero h +end + end expand section char_p From 923ae0bc059a456feeeb08fcbd0749439904777e Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Wed, 4 May 2022 15:53:28 +0000 Subject: [PATCH 104/153] =?UTF-8?q?feat(group=5Ftheory/free=5Fgroup):=20is?= =?UTF-8?q?=5Ffree=5Fgroup=20via=20`free=5Fgroup=20X=20=E2=89=83*=20G`=20(?= =?UTF-8?q?#13633)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The previous definition of the `is_free_group` class was defined via the universal property of free groups, which is intellectually pleasing, but technically annoying, due to the universe problems of quantifying over “all other groups” in the definition. To work around them, many definitions had to be duplicated. This changes the definition of `is_free_group` to contain an isomorphism between the `free_group` over the generator and `G`. It also moves this class into `free_group.lean`, so that it can be found more easily. Relevant Zulip thread: A previous attempt at reforming `is_free_group` to unbundle the set of generators (`is_freely_generated_by G X`) is on branch `joachim/is_freely_generated_by`, but it wasn't very elegant to use in some places. --- src/group_theory/free_product.lean | 19 +- src/group_theory/is_free_group.lean | 240 ++++++++++--------------- src/group_theory/nielsen_schreier.lean | 10 +- 3 files changed, 113 insertions(+), 156 deletions(-) diff --git a/src/group_theory/free_product.lean b/src/group_theory/free_product.lean index b36ecf6efa13a..8471ceda85989 100644 --- a/src/group_theory/free_product.lean +++ b/src/group_theory/free_product.lean @@ -704,15 +704,16 @@ end ping_pong_lemma instance {ι : Type*} (G : ι → Type*) [∀ i, group (G i)] [hG : ∀ i, is_free_group (G i)] : is_free_group (free_product G) := { generators := Σ i, is_free_group.generators (G i), - of := λ x, free_product.of (is_free_group.of x.2), - unique_lift' := - begin - introsI X _ f, - refine ⟨free_product.lift (λ i, is_free_group.lift (λ x, f ⟨i, x⟩)), _ ⟩, - split, - { simp, }, - { intros g hfg, ext i x, simpa using hfg ⟨i, x⟩, } - end, } + mul_equiv := + monoid_hom.to_mul_equiv + (free_group.lift (λ (x : Σ i, is_free_group.generators (G i)), + free_product.of (is_free_group.of x.2 : G x.1))) + (free_product.lift (λ (i : ι), + (is_free_group.lift (λ (x : is_free_group.generators (G i)), + free_group.of (⟨i, x⟩ : Σ i, is_free_group.generators (G i))) + : G i →* (free_group (Σ i, is_free_group.generators (G i)))))) + (by {ext, simp, }) + (by {ext, simp, }) } /-- A free group is a free product of copies of the free_group over one generator. -/ diff --git a/src/group_theory/is_free_group.lean b/src/group_theory/is_free_group.lean index 8a86d700a2247..528ee7e69965d 100644 --- a/src/group_theory/is_free_group.lean +++ b/src/group_theory/is_free_group.lean @@ -1,175 +1,131 @@ /- Copyright (c) 2021 David Wärn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: David Wärn, Eric Wieser +Authors: David Wärn, Eric Wieser, Joachim Breitner -/ import group_theory.free_group /-! # Free groups structures on arbitrary types -This file defines the universal property of free groups, and proves some things about -groups with this property. For an explicit construction of free groups, see -`group_theory/free_group`. +This file defines a type class for type that are free groups, together with the usual operations. +The type class can be instantiated by providing an isomorphim to the canonical free group, or by +proving that the universal property holds. + +For the explicit construction of free groups, see `group_theory/free_group`. ## Main definitions * `is_free_group G` - a typeclass to indicate that `G` is free over some generators -* `is_free_group.lift` - the (noncomputable) universal property of the free group -* `is_free_group.to_free_group` - any free group with generators `A` is equivalent to - `free_group A`. +* `is_free_group.of` - the canonical injection of `G`'s generators into `G` +* `is_free_group.lift` - the universal property of the free group -## Implementation notes +## Main results -While the typeclass only requires the universal property hold within a single universe `u`, our -explicit construction of `free_group` allows this to be extended universe polymorphically. The -primed definition names in this file refer to the non-polymorphic versions. +* `is_free_group.to_free_group` - any free group with generators `A` is equivalent to + `free_group A`. +* `is_free_group.unique_lift` - the universal property of a free group +* `is_free_group.of_unique_lift` - constructing `is_free_group` from the universal property -/ -noncomputable theory -universes u w -/-- `is_free_group G` means that `G` has the universal property of a free group, -That is, it has a family `generators G` of elements, such that a group homomorphism -`G →* X` is uniquely determined by a function `generators G → X`. -/ -class is_free_group (G : Type u) [group G] : Type (u+1) := -(generators : Type u) -(of : generators → G) -(unique_lift' : ∀ {X : Type u} [group X] (f : generators → X), - ∃! F : G →* X, ∀ a, F (of a) = f a) - -instance free_group_is_free_group {A} : is_free_group (free_group A) := -{ generators := A, - of := free_group.of, - unique_lift' := begin - introsI X _ f, - have := free_group.lift.symm.bijective.exists_unique f, - simp_rw function.funext_iff at this, - exact this, - end } +universes u -namespace is_free_group +/-- `is_free_group G` means that `G` isomorphic to a free group. -/ +class is_free_group (G : Type u) [group G] := +(generators : Type u) +(mul_equiv [] : free_group generators ≃* G) -variables {G H : Type u} {X : Type w} [group G] [group H] [group X] [is_free_group G] +instance (X : Type*) : is_free_group (free_group X) := +{ generators := X, + mul_equiv := mul_equiv.refl _ } -/-- The equivalence between functions on the generators and group homomorphisms from a free group -given by those generators. -/ -@[simps symm_apply] -def lift' : (generators G → H) ≃ (G →* H) := -{ to_fun := λ f, classical.some (unique_lift' f), - inv_fun := λ F, F ∘ of, - left_inv := λ f, funext (classical.some_spec (unique_lift' f)).left, - right_inv := λ F, ((classical.some_spec (unique_lift' (F ∘ of))).right F (λ _, rfl)).symm } +namespace is_free_group -@[simp] lemma lift'_of (f : generators G → H) (a : generators G) : (lift' f) (of a) = f a := -congr_fun (lift'.symm_apply_apply f) a +variables (G : Type*) [group G] [is_free_group G] -@[simp] lemma lift'_eq_free_group_lift {A : Type u} : - (@lift' (free_group A) H _ _ _) = free_group.lift := -begin - -- TODO: `apply equiv.symm_bijective.injective`, - rw [←free_group.lift.symm_symm, ←(@lift' (free_group A) H _ _ _).symm_symm], - congr' 1, - ext, - refl, -end - -@[simp] lemma of_eq_free_group_of {A : Type u} : (@of (free_group A) _ _) = free_group.of := -rfl - -@[ext] -lemma ext_hom' ⦃f g : G →* H⦄ (h : ∀ a, f (of a) = g (of a)) : - f = g := -lift'.symm.injective $ funext h - -/-- Being a free group transports across group isomorphisms within a universe. -/ -def of_mul_equiv (h : G ≃* H) : is_free_group H := -{ generators := generators G, - of := h ∘ of, - unique_lift' := begin - introsI X _ f, - refine ⟨(lift' f).comp h.symm.to_monoid_hom, _, _⟩, - { simp }, - intros F' hF', - suffices : F'.comp h.to_monoid_hom = lift' f, - { rw ←this, ext, apply congr_arg, symmetry, apply mul_equiv.apply_symm_apply }, - ext, - simp [hF'], - end } +/-- Any free group is isomorphic to "the" free group. -/ +@[simps] def to_free_group : G ≃* free_group (generators G) := (mul_equiv G).symm -/-! -### Universe-polymorphic definitions +variable {G} +/-- The canonical injection of G's generators into G -/ +def of : generators G → G := (mul_equiv G).to_fun ∘ free_group.of -The primed definitions and lemmas above require `G` and `H` to be in the same universe `u`. -The lemmas below use `X` in a different universe `w` --/ +@[simp] lemma of_eq_free_group_of {A : Type u} : (@of (free_group A) _ _ ) = free_group.of := rfl -variable (G) +variables {H : Type*} [group H] -/-- Any free group is isomorphic to "the" free group. -/ -@[simps] def to_free_group : G ≃* free_group (generators G) := -{ to_fun := lift' free_group.of, - inv_fun := free_group.lift of, - left_inv := - suffices (free_group.lift of).comp (lift' free_group.of) = monoid_hom.id G, - from monoid_hom.congr_fun this, - by { ext, simp }, - right_inv := - suffices - (lift' free_group.of).comp (free_group.lift of) = monoid_hom.id (free_group (generators G)), - from monoid_hom.congr_fun this, - by { ext, simp }, - map_mul' := (lift' free_group.of).map_mul } +/-- The equivalence between functions on the generators and group homomorphisms from a free group +given by those generators. -/ +def lift : (generators G → H) ≃ (G →* H) := +free_group.lift.trans + { to_fun := λ f, f.comp (mul_equiv G).symm.to_monoid_hom, + inv_fun := λ f, f.comp (mul_equiv G).to_monoid_hom, + left_inv := λ f, by { ext, simp, }, + right_inv := λ f, by { ext, simp, }, } -variable {G} +@[simp] lemma lift'_eq_free_group_lift {A : Type u} : + (@lift (free_group A) _ _ H _) = free_group.lift := rfl -private lemma lift_right_inv_aux (F : G →* X) : - free_group.lift.symm (F.comp (to_free_group G).symm.to_monoid_hom) = F ∘ of := -by { ext, simp } - -/-- A universe-polymorphic version of `is_free_group.lift'`. -/ -@[simps symm_apply] -def lift : (generators G → X) ≃ (G →* X) := -{ to_fun := λ f, (free_group.lift f).comp (to_free_group G).to_monoid_hom, - inv_fun := λ F, F ∘ of, - left_inv := λ f, free_group.lift.injective begin - ext x, - simp, - end, - right_inv := λ F, begin - dsimp, - rw ←lift_right_inv_aux, - simp only [equiv.apply_symm_apply], - ext x, - dsimp only [monoid_hom.comp_apply, mul_equiv.coe_to_monoid_hom], - rw mul_equiv.symm_apply_apply, - end} - -@[ext] -lemma ext_hom ⦃f g : G →* X⦄ (h : ∀ a, f (of a) = g (of a)) : - f = g := -is_free_group.lift.symm.injective $ funext h - -@[simp] lemma lift_of (f : generators G → X) (a : generators G) : (lift f) (of a) = f a := +@[simp] lemma lift_of (f : generators G → H) (a : generators G) : lift f (of a) = f a := congr_fun (lift.symm_apply_apply f) a -@[simp] lemma lift_eq_free_group_lift {A : Type u} : - (@lift (free_group A) H _ _ _) = free_group.lift := -begin - -- TODO: `apply equiv.symm_bijective.injective`, - rw [←free_group.lift.symm_symm, ←(@lift (free_group A) H _ _ _).symm_symm], - congr' 1, - ext, - refl, -end - -/-- A universe-polymorphic version of `unique_lift`. -/ -lemma unique_lift {X : Type w} [group X] (f : generators G → X) : - ∃! F : G →* X, ∀ a, F (of a) = f a := -begin - have := lift.symm.bijective.exists_unique f, - simp_rw function.funext_iff at this, - exact this, -end +@[simp] lemma lift_symm_apply (f : G →* H) (a : generators G) : (lift.symm f) a = f (of a) := +rfl + +@[ext] lemma ext_hom ⦃f g : G →* H⦄ (h : ∀ (a : generators G), f (of a) = g (of a)) : f = g := +lift.symm.injective (funext h) + +/-- The universal property of a free group: A functions from the generators of `G` to another +group extends in a unique way to a homomorphism from `G`. + +Note that since `is_free_group.lift` is expressed as a bijection, it already +expresses the universal property. -/ +lemma unique_lift (f : generators G → H) : ∃! F : G →* H, ∀ a, F (of a) = f a := +by simpa only [function.funext_iff] using lift.symm.bijective.exists_unique f + +/-- If a group satisfies the universal property of a free group, then it is a free group, where +the universal property is expressed as in `is_free_group.lift` and its properties. -/ +def of_lift {G : Type u} [group G] (X : Type u) + (of : X → G) + (lift : ∀ {H : Type u} [group H], by exactI (X → H) ≃ (G →* H)) + (lift_of : ∀ {H : Type u} [group H], by exactI ∀ (f : X → H) a, lift f (of a) = f a) : + is_free_group G := +{ generators := X, + mul_equiv := monoid_hom.to_mul_equiv + (free_group.lift of) + (lift free_group.of) + begin + apply free_group.ext_hom, intro x, + simp only [monoid_hom.coe_comp, function.comp_app, monoid_hom.id_apply, free_group.lift.of, + lift_of], + end + begin + let lift_symm_of : ∀ {H : Type u} [group H], by exactI ∀ (f : G →* H) a, + lift.symm f a = f (of a) := by introsI H _ f a; simp [← lift_of (lift.symm f)], + apply lift.symm.injective, ext x, + simp only [monoid_hom.coe_comp, function.comp_app, monoid_hom.id_apply, + free_group.lift.of, lift_of, lift_symm_of], + end } + +/-- If a group satisfies the universal property of a free group, then it is a free group, where +the universal property is expressed as in `is_free_group.unique_lift`. -/ +noncomputable +def of_unique_lift {G : Type u} [group G] (X : Type u) + (of : X → G) + (h : ∀ {H : Type u} [group H] (f : X → H), by exactI ∃! F : G →* H, ∀ a, F (of a) = f a) : + is_free_group G := +let lift {H : Type u} [group H] : by exactI (X → H) ≃ (G →* H) := by exactI + { to_fun := λ f, classical.some (h f), + inv_fun := λ F, F ∘ of, + left_inv := λ f, funext (classical.some_spec (h f)).left, + right_inv := λ F, ((classical.some_spec (h (F ∘ of))).right F (λ _, rfl)).symm } in +let lift_of {H : Type u} [group H] (f : X → H) (a : X) : by exactI lift f (of a) = f a := + by exactI congr_fun (lift.symm_apply_apply f) a in +of_lift X of @lift @lift_of + +/-- Being a free group transports across group isomorphisms. -/ +def of_mul_equiv {H : Type*} [group H] (h : G ≃* H) : is_free_group H := +{ generators := generators G, mul_equiv := (mul_equiv G).trans h } end is_free_group diff --git a/src/group_theory/nielsen_schreier.lean b/src/group_theory/nielsen_schreier.lean index 0ae93548c599d..f9f646ce1c411 100644 --- a/src/group_theory/nielsen_schreier.lean +++ b/src/group_theory/nielsen_schreier.lean @@ -190,10 +190,10 @@ end /-- Given a free groupoid and an arborescence of its generating quiver, the vertex group at the root is freely generated by loops coming from generating arrows in the complement of the tree. -/ -def End_is_free : is_free_group (End (root' T)) := -{ generators := set.compl (wide_subquiver_equiv_set_total $ wide_subquiver_symmetrify T), - of := λ e, loop_of_hom T (of e.val.hom), - unique_lift' := begin +def End_is_free : is_free_group (End (root' T)) := is_free_group.of_unique_lift + (set.compl (wide_subquiver_equiv_set_total $ wide_subquiver_symmetrify T)) + (λ e, loop_of_hom T (of e.val.hom)) + begin introsI X _ f, let f' : labelling (generators G) X := λ a b e, if h : e ∈ wide_subquiver_symmetrify T a b then 1 @@ -226,7 +226,7 @@ def End_is_free : is_free_group (End (root' T)) := split_ifs, { rw [loop_of_hom_eq_id T e h, ←End.one_def, E.map_one] }, { exact hE ⟨⟨a, b, e⟩, h⟩ } } - end } + end end spanning_tree From 60ad8448586c787c4cc409173ef2923810809447 Mon Sep 17 00:00:00 2001 From: tb65536 Date: Wed, 4 May 2022 15:53:29 +0000 Subject: [PATCH 105/153] feat(group_theory/complement): API lemmas relating `range_mem_transversals` and `to_equiv` (#13694) This PR adds an API lemma relating `range_mem_left_transversals` (the main way of constructing left transversals) and `mem_left_transversals.to_equiv` (one of the main constructions from left transversals), and a similar lemma relating the right versions. --- src/group_theory/complement.lean | 15 +++++++++++++++ 1 file changed, 15 insertions(+) diff --git a/src/group_theory/complement.lean b/src/group_theory/complement.lean index 56ab33596f457..bdd887bfe7d56 100644 --- a/src/group_theory/complement.lean +++ b/src/group_theory/complement.lean @@ -241,6 +241,13 @@ noncomputable def to_equiv (hS : S ∈ subgroup.left_transversals (H : set G)) : quotient.mk' (to_equiv hS q : G) = q := (to_equiv hS).symm_apply_apply q +@[to_additive] lemma to_equiv_apply {f : G ⧸ H → G} (hf : ∀ q, (f q : G ⧸ H) = q) (q : G ⧸ H) : + (to_equiv (range_mem_left_transversals hf) q : G) = f q := +begin + refine (subtype.ext_iff.mp _).trans (subtype.coe_mk (f q) ⟨q, rfl⟩), + exact (to_equiv (range_mem_left_transversals hf)).apply_eq_iff_eq_symm_apply.mpr (hf q).symm, +end + /-- A left transversal can be viewed as a function mapping each element of the group to the chosen representative from that left coset. -/ @[to_additive "A left transversal can be viewed as a function mapping each element of the group @@ -270,6 +277,14 @@ noncomputable def to_equiv (hS : S ∈ subgroup.right_transversals (H : set G)) (q : quotient (quotient_group.right_rel H)) : quotient.mk' (to_equiv hS q : G) = q := (to_equiv hS).symm_apply_apply q +@[to_additive] lemma to_equiv_apply {f : quotient (quotient_group.right_rel H) → G} + (hf : ∀ q, quotient.mk' (f q) = q) (q : quotient (quotient_group.right_rel H)) : + (to_equiv (range_mem_right_transversals hf) q : G) = f q := +begin + refine (subtype.ext_iff.mp _).trans (subtype.coe_mk (f q) ⟨q, rfl⟩), + exact (to_equiv (range_mem_right_transversals hf)).apply_eq_iff_eq_symm_apply.mpr (hf q).symm, +end + /-- A right transversal can be viewed as a function mapping each element of the group to the chosen representative from that right coset. -/ @[to_additive "A right transversal can be viewed as a function mapping each element of the group From 46023701ec7ed21d8f4126834a758b54e826386b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 4 May 2022 15:53:30 +0000 Subject: [PATCH 106/153] feat(set_theory/game/birthday): More basic lemmas on birthdays (#13729) --- src/set_theory/game/birthday.lean | 13 +++++++++++++ src/set_theory/game/pgame.lean | 2 ++ 2 files changed, 15 insertions(+) diff --git a/src/set_theory/game/birthday.lean b/src/set_theory/game/birthday.lean index db9434bf2d621..673f4f16d2c49 100644 --- a/src/set_theory/game/birthday.lean +++ b/src/set_theory/game/birthday.lean @@ -88,6 +88,12 @@ theorem relabelling.birthday_congr : ∀ {x y : pgame.{u}}, relabelling x y → end using_well_founded { dec_tac := pgame_wf_tac } +@[simp] theorem birthday_add_zero (x : pgame) : birthday (x + 0) = birthday x := +(add_zero_relabelling x).birthday_congr + +@[simp] theorem birthday_zero_add (x : pgame) : birthday (0 + x) = birthday x := +(zero_add_relabelling x).birthday_congr + @[simp] theorem birthday_eq_zero (x : pgame) : birthday x = 0 ↔ is_empty x.left_moves ∧ is_empty x.right_moves := by rw [birthday_def, ordinal.max_eq_zero, ordinal.lsub_eq_zero_iff, ordinal.lsub_eq_zero_iff] @@ -95,4 +101,11 @@ by rw [birthday_def, ordinal.max_eq_zero, ordinal.lsub_eq_zero_iff, ordinal.lsub @[simp] theorem birthday_zero : birthday 0 = 0 := 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], + exact max_bot_right 1 +end + end pgame diff --git a/src/set_theory/game/pgame.lean b/src/set_theory/game/pgame.lean index e423d78dd20dc..7fbfbb9587a4b 100644 --- a/src/set_theory/game/pgame.lean +++ b/src/set_theory/game/pgame.lean @@ -840,6 +840,8 @@ end instance : has_add pgame := ⟨add⟩ +@[simp] theorem nat_one : ((1 : ℕ) : pgame) = 0 + 1 := rfl + /-- `x + 0` has exactly the same moves as `x`. -/ def add_zero_relabelling : Π (x : pgame.{u}), relabelling (x + 0) x | (mk xl xr xL xR) := From 0038a0436f689d37fd3ad00d69d617c338ae6028 Mon Sep 17 00:00:00 2001 From: Bhavik Mehta Date: Wed, 4 May 2022 15:53:31 +0000 Subject: [PATCH 107/153] feat(data/int/cast): int cast division lemmas (#13929) Adds lemmas for passing int cast through division, and renames the nat versions from `nat.cast_dvd` to `nat.cast_div`. Also some golf. Co-authored-by: Scott Morrison --- src/algebra/char_zero.lean | 9 ++++----- src/data/int/cast.lean | 8 ++++++++ src/data/int/char_zero.lean | 9 +++++++++ src/data/nat/cast.lean | 5 ++--- src/data/nat/totient.lean | 2 +- src/number_theory/bernoulli.lean | 6 +++--- src/number_theory/von_mangoldt.lean | 2 +- src/ring_theory/discriminant.lean | 2 +- src/ring_theory/power_series/well_known.lean | 2 +- 9 files changed, 30 insertions(+), 15 deletions(-) diff --git a/src/algebra/char_zero.lean b/src/algebra/char_zero.lean index b99b35254363e..86887715e33d7 100644 --- a/src/algebra/char_zero.lean +++ b/src/algebra/char_zero.lean @@ -101,13 +101,12 @@ lemma cast_add_one_ne_zero (n : ℕ) : (n + 1 : R) ≠ 0 := by exact_mod_cast n.succ_ne_zero @[simp, norm_cast] -theorem cast_dvd_char_zero {k : Type*} [field k] [char_zero k] {m n : ℕ} +theorem cast_div_char_zero {k : Type*} [field k] [char_zero k] {m n : ℕ} (n_dvd : n ∣ m) : ((m / n : ℕ) : k) = m / n := begin - by_cases hn : n = 0, - { subst hn, - simp }, - { exact cast_dvd n_dvd (cast_ne_zero.mpr hn), }, + rcases eq_or_ne n 0 with rfl | hn, + { simp }, + { exact cast_div n_dvd (cast_ne_zero.2 hn), }, end end nat diff --git a/src/data/int/cast.lean b/src/data/int/cast.lean index 0d5b983deb0cc..8a0be11ecfbf3 100644 --- a/src/data/int/cast.lean +++ b/src/data/int/cast.lean @@ -111,6 +111,14 @@ by simp [sub_eq_add_neg] | -[1+ m] -[1+ n] := show (((m + 1) * (n + 1) : ℕ) : α) = -(m + 1) * -(n + 1), by rw [nat.cast_mul, nat.cast_add_one, nat.cast_add_one, neg_mul_neg] +@[simp] theorem cast_div [field α] {m n : ℤ} (n_dvd : n ∣ m) (n_nonzero : (n : α) ≠ 0) : + ((m / n : ℤ) : α) = m / n := +begin + rcases n_dvd with ⟨k, rfl⟩, + have : n ≠ 0, { rintro rfl, simpa using n_nonzero }, + rw [int.mul_div_cancel_left _ this, int.cast_mul, mul_div_cancel_left _ n_nonzero], +end + /-- `coe : ℤ → α` as an `add_monoid_hom`. -/ def cast_add_hom (α : Type*) [add_group α] [has_one α] : ℤ →+ α := ⟨coe, cast_zero, cast_add⟩ diff --git a/src/data/int/char_zero.lean b/src/data/int/char_zero.lean index 0052737ae63d0..f8209451f13cb 100644 --- a/src/data/int/char_zero.lean +++ b/src/data/int/char_zero.lean @@ -35,6 +35,15 @@ theorem cast_injective [add_group α] [has_one α] [char_zero α] : function.inj theorem cast_ne_zero [add_group α] [has_one α] [char_zero α] {n : ℤ} : (n : α) ≠ 0 ↔ n ≠ 0 := not_congr cast_eq_zero +@[simp, norm_cast] +theorem cast_div_char_zero {k : Type*} [field k] [char_zero k] {m n : ℤ} + (n_dvd : n ∣ m) : ((m / n : ℤ) : k) = m / n := +begin + rcases eq_or_ne n 0 with rfl | hn, + { simp [int.div_zero] }, + { exact cast_div n_dvd (cast_ne_zero.mpr hn), }, +end + end int lemma ring_hom.injective_int {α : Type*} [ring α] (f : ℤ →+* α) [char_zero α] : diff --git a/src/data/nat/cast.lean b/src/data/nat/cast.lean index 2860fad6544e8..e9cd6999cfd12 100644 --- a/src/data/nat/cast.lean +++ b/src/data/nat/cast.lean @@ -132,13 +132,12 @@ eq_sub_of_add_eq $ by rw [← cast_add, tsub_add_cancel_of_le h] | (n+1) := (cast_add _ _).trans $ show ((m * n : ℕ) : α) + m = m * (n + 1), by rw [cast_mul n, left_distrib, mul_one] -@[simp] theorem cast_dvd {α : Type*} [field α] {m n : ℕ} (n_dvd : n ∣ m) (n_nonzero : (n:α) ≠ 0) : +@[simp] theorem cast_div [field α] {m n : ℕ} (n_dvd : n ∣ m) (n_nonzero : (n : α) ≠ 0) : ((m / n : ℕ) : α) = m / n := begin rcases n_dvd with ⟨k, rfl⟩, have : n ≠ 0, {rintro rfl, simpa using n_nonzero}, - rw nat.mul_div_cancel_left _ (pos_iff_ne_zero.2 this), - rw [nat.cast_mul, mul_div_cancel_left _ n_nonzero], + rw [nat.mul_div_cancel_left _ this.bot_lt, cast_mul, mul_div_cancel_left _ n_nonzero], end /-- `coe : ℕ → α` as a `ring_hom` -/ diff --git a/src/data/nat/totient.lean b/src/data/nat/totient.lean index 3ba8b8d2732e3..0896597ca2290 100644 --- a/src/data/nat/totient.lean +++ b/src/data/nat/totient.lean @@ -335,7 +335,7 @@ begin { rw [←cast_prod, cast_ne_zero, ←zero_lt_iff, ←prod_factorization_eq_prod_factors], exact prod_pos (λ p hp, pos_of_mem_factorization hp) }, simp only [totient_eq_div_factors_mul n, prod_prime_factors_dvd n, cast_mul, cast_prod, - cast_dvd_char_zero, mul_comm_div', mul_right_inj' hn', div_eq_iff hpQ, ←prod_mul_distrib], + cast_div_char_zero, mul_comm_div', mul_right_inj' hn', div_eq_iff hpQ, ←prod_mul_distrib], refine prod_congr rfl (λ p hp, _), have hp := pos_of_mem_factors (list.mem_to_finset.mp hp), have hp' : (p : ℚ) ≠ 0 := cast_ne_zero.mpr hp.ne.symm, diff --git a/src/number_theory/bernoulli.lean b/src/number_theory/bernoulli.lean index 37adaf8e99cc0..68e1350fd933d 100644 --- a/src/number_theory/bernoulli.lean +++ b/src/number_theory/bernoulli.lean @@ -148,7 +148,7 @@ begin have := factorial_mul_factorial_dvd_factorial_add i j, field_simp [mul_comm _ (bernoulli' i), mul_assoc, add_choose], rw_mod_cast [mul_comm (j + 1), mul_div_assoc, ← mul_assoc], - rw [cast_mul, cast_mul, mul_div_mul_right, cast_dvd_char_zero, cast_mul], + rw [cast_mul, cast_mul, mul_div_mul_right, cast_div_char_zero, cast_mul], assumption', end @@ -256,7 +256,7 @@ begin have hj : (j.succ : ℚ) ≠ 0 := by exact_mod_cast succ_ne_zero j, field_simp [← h, mul_ne_zero hj (hfact j), hfact i, mul_comm _ (bernoulli i), mul_assoc], rw_mod_cast [mul_comm (j + 1), mul_div_assoc, ← mul_assoc], - rw [cast_mul, cast_mul, mul_div_mul_right _ _ hj, add_choose, cast_dvd_char_zero], + rw [cast_mul, cast_mul, mul_div_mul_right _ _ hj, add_choose, cast_div_char_zero], apply factorial_mul_factorial_dvd_factorial_add, end @@ -287,7 +287,7 @@ begin rw [choose_eq_factorial_div_factorial h.le, eq_comm, div_eq_iff (hne q.succ), succ_eq_add_one, mul_assoc _ _ ↑q.succ!, mul_comm _ ↑q.succ!, ← mul_assoc, div_mul_eq_mul_div, mul_comm (↑n ^ (q - m + 1)), ← mul_assoc _ _ (↑n ^ (q - m + 1)), ← one_div, mul_one_div, - div_div_eq_div_mul, tsub_add_eq_add_tsub (le_of_lt_succ h), cast_dvd, cast_mul], + div_div_eq_div_mul, tsub_add_eq_add_tsub (le_of_lt_succ h), cast_div, cast_mul], { ring }, { exact factorial_mul_factorial_dvd_factorial h.le }, { simp [hne] } }, diff --git a/src/number_theory/von_mangoldt.lean b/src/number_theory/von_mangoldt.lean index e0fa12b38626c..6f96ac672b62f 100644 --- a/src/number_theory/von_mangoldt.lean +++ b/src/number_theory/von_mangoldt.lean @@ -124,7 +124,7 @@ begin { rw [cast_ne_zero], rintro rfl, exact hn (by simpa using mn) }, - rw [nat.cast_dvd mn this, real.log_div (cast_ne_zero.2 hn) this, neg_sub, mul_sub] }, + rw [nat.cast_div mn this, real.log_div (cast_ne_zero.2 hn) this, neg_sub, mul_sub] }, rw [this, sum_sub_distrib, ←sum_mul, ←int.cast_sum, ←coe_mul_zeta_apply, eq_comm, sub_eq_self, moebius_mul_coe_zeta, mul_eq_zero, int.cast_eq_zero], rcases eq_or_ne n 1 with hn | hn; diff --git a/src/ring_theory/discriminant.lean b/src/ring_theory/discriminant.lean index 9036d18c3b99b..cb55318b03c98 100644 --- a/src/ring_theory/discriminant.lean +++ b/src/ring_theory/discriminant.lean @@ -203,7 +203,7 @@ begin have hle : 1 ≤ pb.dim, { rw [← hn, nat.one_le_iff_ne_zero, ← zero_lt_iff, finite_dimensional.finrank_pos_iff], apply_instance }, - rw [hn, nat.cast_dvd h₂ hne, nat.cast_mul, nat.cast_sub hle], + rw [hn, nat.cast_div h₂ hne, nat.cast_mul, nat.cast_sub hle], field_simp, ring, end diff --git a/src/ring_theory/power_series/well_known.lean b/src/ring_theory/power_series/well_known.lean index 12f8be453a4fd..0e1ee81f3f6fd 100644 --- a/src/ring_theory/power_series/well_known.lean +++ b/src/ring_theory/power_series/well_known.lean @@ -110,7 +110,7 @@ begin symmetry, rw [div_eq_iff, div_mul_eq_mul_div, one_mul, choose_eq_factorial_div_factorial], norm_cast, - rw cast_dvd_char_zero, + rw cast_div_char_zero, { apply factorial_mul_factorial_dvd_factorial (mem_range_succ_iff.1 hx), }, { apply mem_range_succ_iff.1 hx, }, { rintros h, apply factorial_ne_zero n, rw cast_eq_zero.1 h, }, From abcd601cd12c86fce3b85942817e1ca78e31b0c7 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Wed, 4 May 2022 17:58:24 +0000 Subject: [PATCH 108/153] =?UTF-8?q?fix(src/tactic/alias):=20Teach=20`get?= =?UTF-8?q?=5Falias=5Ftarget`=20about=20`alias=20f=20=E2=86=94=20a=20b`=20?= =?UTF-8?q?(#13742)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit the `get_alias_target` function in `alias.lean` is used by the `to_additive` command to add the “Alias of …” docstring when creating an additive version of an existing alias (this was #13330). But `get_alias_target` did not work for `alias f ↔ a b`. This fixes it by extending the `alias_attr` map to not just store whether a defintion is an alias, but also what it is an alias of. Much more principled than trying to reconstruct the alias target from the RHS of the alias definition. Note that `alias` currently says “Alias of `foo_iff`” even though it’s really an alias of `foo_iff.mp`. This is an existing bug, not fixed in this PR – the effect is just that this “bug” will uniformly apply to additive lemmas as well. Hopefully will get rid of plenty of nolint.txt entries, and create better docs. Also improve the test file for the linter significantly. --- src/algebra/parity.lean | 10 ++--- src/tactic/alias.lean | 15 +++---- test/lint_to_additive_doc.lean | 75 +++++++++++++++++----------------- 3 files changed, 47 insertions(+), 53 deletions(-) diff --git a/src/algebra/parity.lean b/src/algebra/parity.lean index 048cac9bdf775..0e2077bf429fc 100644 --- a/src/algebra/parity.lean +++ b/src/algebra/parity.lean @@ -53,13 +53,11 @@ by simp [is_square, pow_two] alias is_square_iff_exists_sq ↔ is_square.exists_sq is_square_of_exists_sq -attribute [to_additive even.exists_two_nsmul] is_square.exists_sq -/-- Alias of the forwards direction of `even_iff_exists_two_nsmul`. -/ -add_decl_doc even.exists_two_nsmul +attribute [to_additive even.exists_two_nsmul "Alias of the forwards direction of +`even_iff_exists_two_nsmul`."] is_square.exists_sq -attribute [to_additive even_of_exists_two_nsmul] is_square_of_exists_sq -/-- Alias of the backwards direction of `even_iff_exists_two_nsmul`. -/ -add_decl_doc even_of_exists_two_nsmul +attribute [to_additive even_of_exists_two_nsmul "Alias of the backwards direction of +`even_iff_exists_two_nsmul`."] is_square_of_exists_sq @[simp, to_additive] lemma is_square_one [mul_one_class α] : is_square (1 : α) := ⟨1, (mul_one _).symm⟩ diff --git a/src/tactic/alias.lean b/src/tactic/alias.lean index 1353c9fac8798..a48ebe2afc3b6 100644 --- a/src/tactic/alias.lean +++ b/src/tactic/alias.lean @@ -46,7 +46,7 @@ open lean.parser tactic interactive namespace tactic.alias -@[user_attribute] meta def alias_attr : user_attribute := +@[user_attribute] meta def alias_attr : user_attribute unit name := { name := `alias, descr := "This definition is an alias of another.", parser := failed } meta def alias_direct (d : declaration) (doc : string) (al : name) : tactic unit := @@ -59,7 +59,7 @@ do updateex_env $ λ env, declaration.thm al ls t $ task.pure $ expr.const n (level.param <$> ls) | _ := undefined end), - alias_attr.set al () tt, + alias_attr.set al d.to_name tt, add_doc_string al doc meta def mk_iff_mp_app (iffmp : name) : expr → (ℕ → expr) → tactic expr @@ -74,7 +74,7 @@ meta def alias_iff (d : declaration) (doc : string) (al : name) (iffmp : name) : 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 () tt, + alias_attr.set al d.to_name tt, add_doc_string al doc meta def make_left_right : name → tactic (name × name) @@ -156,12 +156,7 @@ meta def get_lambda_body : expr → expr meta def get_alias_target (n : name) : tactic (option name) := do tt ← has_attribute' `alias n | pure none, - d ← get_decl n, - let (head, args) := (get_lambda_body d.value).get_app_fn_args, - let head := if head.is_constant_of `iff.mp ∨ head.is_constant_of `iff.mpr then - expr.get_app_fn (head.ith_arg 2) - else head, - guardb $ head.is_constant, - pure $ head.const_name + v ← alias_attr.get_param n, + pure $ some v end tactic.alias diff --git a/test/lint_to_additive_doc.lean b/test/lint_to_additive_doc.lean index b61731f83fd43..51e9f6175b14e 100644 --- a/test/lint_to_additive_doc.lean +++ b/test/lint_to_additive_doc.lean @@ -1,61 +1,62 @@ import algebra.group.to_additive import tactic.alias +/-- Test assertion helpers -/ +meta def assert_ok (n : name) := do + decl ← tactic.get_decl n, + some msg ← linter.to_additive_doc.test decl | pure (), + fail! "Linting {n} complained unexpectedly:\n{msg}" + +meta def assert_complain (n : name) := do + decl ← tactic.get_decl n, + none ← linter.to_additive_doc.test decl | pure (), + fail! "Linting {n} succeeded unexpectedly" + /-- Docstring -/ @[to_additive add_foo] def foo (α : Type*) [has_one α] : α := 1 +run_cmd assert_complain ``foo +run_cmd assert_ok ``add_foo + @[to_additive add_bar "docstring"] def bar (α : Type*) [has_one α] : α := 1 +run_cmd assert_complain ``bar +run_cmd assert_ok ``add_bar + /-- Docstring -/ @[to_additive add_baz "docstring"] def baz (α : Type*) [has_one α] : α := 1 +run_cmd assert_ok ``baz +run_cmd assert_ok ``add_baz + @[to_additive add_quux] def quux (α : Type*) [has_one α] : α := 1 -def no_to_additive (α : Type*) [has_one α] : α := 1 +run_cmd assert_ok ``quux +run_cmd assert_ok ``add_quux + +def without_to_additive (α : Type*) [has_one α] : α := 1 + +run_cmd assert_ok ``without_to_additive -- Aliases always have docstrings, so we do not want to complain if their -- additive version do not alias quux <- quux_alias attribute [to_additive add_quux_alias] quux_alias +run_cmd assert_ok ``quux_alias +run_cmd assert_ok ``add_quux_alias + +-- Same for iff aliases +def a_one_iff_b_one (α : Type*) [has_one α] (a b : α) (h : a = b) : (a = 1) ↔ (b = 1) := by {subst h} +alias a_one_iff_b_one ↔ b_one_of_a_one a_one_of_b_one +attribute [to_additive] a_one_iff_b_one +attribute [to_additive] b_one_of_a_one +attribute [to_additive] a_one_of_b_one -open tactic -run_cmd do - decl ← get_decl ``foo, - res ← linter.to_additive_doc.test decl, - -- linter complains - guard $ res.is_some - -run_cmd do - decl ← get_decl ``bar, - res ← linter.to_additive_doc.test decl, - -- linter complains - guard $ res.is_some - -run_cmd do - decl ← get_decl ``baz, - res ← linter.to_additive_doc.test decl, - -- linter is happy - guard $ res.is_none - -run_cmd do - decl ← get_decl ``quux, - res ← linter.to_additive_doc.test decl, - -- linter is happy - guard $ res.is_none - -run_cmd do - decl ← get_decl ``no_to_additive, - res ← linter.to_additive_doc.test decl, - -- linter is happy - guard $ res.is_none - -run_cmd do - decl ← get_decl ``quux_alias, - res ← linter.to_additive_doc.test decl, - -- linter is happy - guard $ res.is_none +run_cmd assert_ok ``a_one_iff_b_one +run_cmd assert_ok ``a_one_of_b_one +run_cmd assert_ok ``b_one_of_a_one From dd4590af908d6443402aed7fd8ad899917b2a518 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 4 May 2022 17:58:26 +0000 Subject: [PATCH 109/153] refactor(algebra/restrict_scalars): remove global instance on module_orig (#13759) The global instance was conceptually wrong, unnecessary (after avoiding a hack in algebra/lie/base_change.lean), and wreaking havoc in #13713. Co-authored-by: Scott Morrison Co-authored-by: Eric Wieser --- src/algebra/algebra/restrict_scalars.lean | 90 ++++++++++++++++------- src/algebra/lie/base_change.lean | 5 +- src/analysis/normed_space/basic.lean | 14 +++- src/analysis/normed_space/extend.lean | 4 +- src/ring_theory/algebra_tower.lean | 7 -- 5 files changed, 79 insertions(+), 41 deletions(-) diff --git a/src/algebra/algebra/restrict_scalars.lean b/src/algebra/algebra/restrict_scalars.lean index 359974a5458e6..03f6a82c1bb31 100644 --- a/src/algebra/algebra/restrict_scalars.lean +++ b/src/algebra/algebra/restrict_scalars.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Yury Kudryashov -/ -import algebra.algebra.basic +import algebra.algebra.tower /-! @@ -16,10 +16,14 @@ typeclass instead. ## Main definitions * `restrict_scalars R S M`: the `S`-module `M` viewed as an `R` module when `S` is an `R`-algebra. -* `restrict_scalars.linear_equiv : restrict_scalars R S M ≃ₗ[S] M`: the equivalence as an - `S`-module between the restricted and origianl space. -* `restrict_scalars.alg_equiv : restrict_scalars R S A ≃ₐ[S] A`: the equivalence as an `S`-algebra - between the restricted and original space. + Note that by default we do *not* have a `module S (restrict_scalars R S M)` instance + for the original action. + This is available as a def `restrict_scalars.module_orig` if really needed. +* `restrict_scalars.add_equiv : restrict_scalars R S M ≃+ M`: the additive equivalence + between the restricted and original space (in fact, they are definitionally equal, + but sometimes it is helpful to avoid using this fact, to keep instances from leaking). +* `restrict_scalars.ring_equiv : restrict_scalars R S A ≃+* A`: the ring equivalence + between the restricted and original space when the module is an algebra. ## See also @@ -76,16 +80,19 @@ instance [I : add_comm_monoid M] : add_comm_monoid (restrict_scalars R S M) := I instance [I : add_comm_group M] : add_comm_group (restrict_scalars R S M) := I -instance restrict_scalars.module_orig [semiring S] [add_comm_monoid M] [I : module S M] : +section module + +section +variables [semiring S] [add_comm_monoid M] + +/-- We temporarily install an action of the original ring on `restrict_sclars R S M`. -/ +def restrict_scalars.module_orig [I : module S M] : module S (restrict_scalars R S M) := I -/-- `restrict_scalars.linear_equiv` is an equivalence of modules over the semiring `S`. -/ -def restrict_scalars.linear_equiv [semiring S] [add_comm_monoid M] [module S M] : - restrict_scalars R S M ≃ₗ[S] M := -linear_equiv.refl S M +variables [comm_semiring R] [algebra R S] [module S M] -section module -variables [semiring S] [add_comm_monoid M] [comm_semiring R] [algebra R S] [module S M] +section +local attribute [instance] restrict_scalars.module_orig /-- When `M` is a module over a ring `S`, and `S` is an algebra over `R`, then `M` inherits a @@ -96,17 +103,44 @@ The preferred way of setting this up is `[module R M] [module S M] [is_scalar_to instance : module R (restrict_scalars R S M) := module.comp_hom M (algebra_map R S) +/-- +This instance is only relevant when `restrict_scalars.module_orig` is available as an instance. +-/ +instance : is_scalar_tower R S (restrict_scalars R S M) := +⟨λ r S M, by { rw [algebra.smul_def, mul_smul], refl }⟩ + +end + +/-- +The `R`-algebra homomorphism from the original coefficient algebra `S` to endomorphisms +of `restrict_scalars R S M`. +-/ +def restrict_scalars.lsmul : S →ₐ[R] module.End R (restrict_scalars R S M) := +begin + -- We use `restrict_scalars.module_orig` in the implementation, + -- but not in the type. + letI : module S (restrict_scalars R S M) := restrict_scalars.module_orig R S M, + exact algebra.lsmul R (restrict_scalars R S M), +end + +end + +variables [add_comm_monoid M] + +/-- `restrict_scalars.add_equiv` is the additive equivalence with the original module. -/ +@[simps] def restrict_scalars.add_equiv : restrict_scalars R S M ≃+ M := +add_equiv.refl M + +variables [comm_semiring R] [semiring S] [algebra R S] [module S M] + lemma restrict_scalars_smul_def (c : R) (x : restrict_scalars R S M) : c • x = ((algebra_map R S c) • x : M) := rfl -@[simp] lemma restrict_scalars.linear_equiv_map_smul (t : R) (x : restrict_scalars R S M) : - restrict_scalars.linear_equiv R S M (t • x) - = (algebra_map R S t) • restrict_scalars.linear_equiv R S M x := +@[simp] lemma restrict_scalars.add_equiv_map_smul (t : R) (x : restrict_scalars R S M) : + restrict_scalars.add_equiv R S M (t • x) + = (algebra_map R S t) • restrict_scalars.add_equiv R S M x := rfl -instance : is_scalar_tower R S (restrict_scalars R S M) := -⟨λ r S M, by { rw [algebra.smul_def, mul_smul], refl }⟩ - end module section algebra @@ -116,16 +150,17 @@ instance [I : ring A] : ring (restrict_scalars R S A) := I instance [I : comm_semiring A] : comm_semiring (restrict_scalars R S A) := I instance [I : comm_ring A] : comm_ring (restrict_scalars R S A) := I -variables [comm_semiring S] [semiring A] - -instance restrict_scalars.algebra_orig [I : algebra S A] : algebra S (restrict_scalars R S A) := I +variables [semiring A] -variables [algebra S A] +/-- Tautological ring isomorphism `restrict_scalars R S A ≃+* A`. -/ +def restrict_scalars.ring_equiv : restrict_scalars R S A ≃+* A := ring_equiv.refl _ -/-- Tautological `S`-algebra isomorphism `restrict_scalars R S A ≃ₐ[S] A`. -/ -def restrict_scalars.alg_equiv : restrict_scalars R S A ≃ₐ[S] A := alg_equiv.refl +variables [comm_semiring S] [algebra S A] [comm_semiring R] [algebra R S] -variables [comm_semiring R] [algebra R S] +@[simp] lemma restrict_scalars.ring_equiv_map_smul (r : R) (x : restrict_scalars R S A) : + restrict_scalars.ring_equiv R S A (r • x) + = (algebra_map R S r) • restrict_scalars.ring_equiv R S A x := +rfl /-- `R ⟶ S` induces `S-Alg ⥤ R-Alg` -/ instance : algebra R (restrict_scalars R S A) := @@ -134,4 +169,9 @@ instance : algebra R (restrict_scalars R S A) := smul_def' := λ _ _, algebra.smul_def _ _, .. (algebra_map S A).comp (algebra_map R S) } +@[simp] lemma restrict_scalars.ring_equiv_algebra_map (r : R) : + restrict_scalars.ring_equiv R S A (algebra_map R (restrict_scalars R S A) r) = + algebra_map S A (algebra_map R S r) := +rfl + end algebra diff --git a/src/algebra/lie/base_change.lean b/src/algebra/lie/base_change.lean index 2f4ae5df7b53d..2403b345978e8 100644 --- a/src/algebra/lie/base_change.lean +++ b/src/algebra/lie/base_change.lean @@ -141,10 +141,9 @@ instance : lie_ring (restrict_scalars R A L) := h variables [comm_ring A] [lie_algebra A L] -@[nolint unused_arguments] instance lie_algebra [comm_ring R] [algebra R A] : lie_algebra R (restrict_scalars R A L) := -{ lie_smul := λ t x y, (lie_smul _ (show L, from x) (show L, from y) : _), - .. (by apply_instance : module R (restrict_scalars R A L)), } +{ lie_smul := λ t x y, (lie_smul (algebra_map R A t) + (restrict_scalars.add_equiv R A L x) (restrict_scalars.add_equiv R A L y) : _) } end restrict_scalars diff --git a/src/analysis/normed_space/basic.lean b/src/analysis/normed_space/basic.lean index 48674e5d1c3fb..89306fcaeb9f4 100644 --- a/src/analysis/normed_space/basic.lean +++ b/src/analysis/normed_space/basic.lean @@ -474,11 +474,17 @@ instance {𝕜 : Type*} {𝕜' : Type*} {E : Type*} [I : semi_normed_group E] : instance {𝕜 : Type*} {𝕜' : Type*} {E : Type*} [I : normed_group E] : normed_group (restrict_scalars 𝕜 𝕜' E) := I -instance module.restrict_scalars.normed_space_orig {𝕜 : Type*} {𝕜' : Type*} {E : Type*} - [normed_field 𝕜'] [semi_normed_group E] [I : normed_space 𝕜' E] : - normed_space 𝕜' (restrict_scalars 𝕜 𝕜' E) := I - instance : normed_space 𝕜 (restrict_scalars 𝕜 𝕜' E) := (normed_space.restrict_scalars 𝕜 𝕜' E : normed_space 𝕜 E) +/-- +The action of the original normed_field on `restrict_scalars 𝕜 𝕜' E`. +This is not an instance as it would be contrary to the purpose of `restrict_scalars`. +-/ +-- If you think you need this, consider instead reproducing `restrict_scalars.lsmul` +-- appropriately modified here. +def module.restrict_scalars.normed_space_orig {𝕜 : Type*} {𝕜' : Type*} {E : Type*} + [normed_field 𝕜'] [semi_normed_group E] [I : normed_space 𝕜' E] : + normed_space 𝕜' (restrict_scalars 𝕜 𝕜' E) := I + end restrict_scalars diff --git a/src/analysis/normed_space/extend.lean b/src/analysis/normed_space/extend.lean index 34bfaeb202b44..fd9f11ecc7330 100644 --- a/src/analysis/normed_space/extend.lean +++ b/src/analysis/normed_space/extend.lean @@ -132,7 +132,7 @@ noncomputable def linear_map.extend_to_𝕜 (fr : (restrict_scalars ℝ 𝕜 F) fr.extend_to_𝕜' lemma linear_map.extend_to_𝕜_apply (fr : (restrict_scalars ℝ 𝕜 F) →ₗ[ℝ] ℝ) (x : F) : - fr.extend_to_𝕜 x = (fr x : 𝕜) - (I : 𝕜) * fr ((I : 𝕜) • x) := rfl + fr.extend_to_𝕜 x = (fr x : 𝕜) - (I : 𝕜) * fr ((I : 𝕜) • x : _) := rfl /-- Extend `fr : restrict_scalars ℝ 𝕜 F →L[ℝ] ℝ` to `F →L[𝕜] 𝕜`. -/ noncomputable def continuous_linear_map.extend_to_𝕜 (fr : (restrict_scalars ℝ 𝕜 F) →L[ℝ] ℝ) : @@ -140,4 +140,4 @@ noncomputable def continuous_linear_map.extend_to_𝕜 (fr : (restrict_scalars fr.extend_to_𝕜' lemma continuous_linear_map.extend_to_𝕜_apply (fr : (restrict_scalars ℝ 𝕜 F) →L[ℝ] ℝ) (x : F) : - fr.extend_to_𝕜 x = (fr x : 𝕜) - (I : 𝕜) * fr ((I : 𝕜) • x) := rfl + fr.extend_to_𝕜 x = (fr x : 𝕜) - (I : 𝕜) * fr ((I : 𝕜) • x : _) := rfl diff --git a/src/ring_theory/algebra_tower.lean b/src/ring_theory/algebra_tower.lean index 08200852bcf18..2253f528d4f21 100644 --- a/src/ring_theory/algebra_tower.lean +++ b/src/ring_theory/algebra_tower.lean @@ -65,13 +65,6 @@ end is_scalar_tower namespace algebra -theorem adjoin_algebra_map' {R : Type u} {S : Type v} {A : Type w} - [comm_semiring R] [comm_semiring S] [semiring A] [algebra R S] [algebra S A] (s : set S) : - adjoin R (algebra_map S (restrict_scalars R S A) '' s) = (adjoin R s).map - ((algebra.of_id S (restrict_scalars R S A)).restrict_scalars R) := -le_antisymm (adjoin_le $ set.image_subset_iff.2 $ λ y hy, ⟨y, subset_adjoin hy, rfl⟩) - (subalgebra.map_le.2 $ adjoin_le $ λ y hy, subset_adjoin ⟨y, hy, rfl⟩) - theorem adjoin_algebra_map (R : Type u) (S : Type v) (A : Type w) [comm_semiring R] [comm_semiring S] [semiring A] [algebra R S] [algebra S A] [algebra R A] [is_scalar_tower R S A] (s : set S) : From 0009ffbaae1afb5c9ee0a8cc1999ff95184870b2 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 4 May 2022 17:58:27 +0000 Subject: [PATCH 110/153] refactor(linear_algebra/charpoly): split file to reduce imports (#13778) While working on representation theory I was annoyed to find that essentially all of field theory was being transitively imported (causing lots of unnecessary recompilation). This improves the situation slightly. Co-authored-by: Scott Morrison --- .../83_friendship_graphs.lean | 2 +- src/linear_algebra/charpoly/basic.lean | 2 + src/linear_algebra/matrix/charpoly/coeff.lean | 81 +------------------ .../matrix/charpoly/finite_field.lean | 60 ++++++++++++++ .../matrix/charpoly/minpoly.lean | 61 ++++++++++++++ src/ring_theory/norm.lean | 2 +- src/ring_theory/trace.lean | 2 +- 7 files changed, 127 insertions(+), 83 deletions(-) create mode 100644 src/linear_algebra/matrix/charpoly/finite_field.lean create mode 100644 src/linear_algebra/matrix/charpoly/minpoly.lean diff --git a/archive/100-theorems-list/83_friendship_graphs.lean b/archive/100-theorems-list/83_friendship_graphs.lean index 30ecea0fe0307..0e4f03120b234 100644 --- a/archive/100-theorems-list/83_friendship_graphs.lean +++ b/archive/100-theorems-list/83_friendship_graphs.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson, Jalex Stark, Kyle Miller -/ import combinatorics.simple_graph.adj_matrix -import linear_algebra.matrix.charpoly.coeff +import linear_algebra.matrix.charpoly.finite_field import data.int.modeq import data.zmod.basic import tactic.interval_cases diff --git a/src/linear_algebra/charpoly/basic.lean b/src/linear_algebra/charpoly/basic.lean index d4bc48e3ffe45..630f07ee36211 100644 --- a/src/linear_algebra/charpoly/basic.lean +++ b/src/linear_algebra/charpoly/basic.lean @@ -6,6 +6,7 @@ Authors: Riccardo Brasca import linear_algebra.free_module.finite.basic import linear_algebra.matrix.charpoly.coeff +import field_theory.minpoly /-! @@ -30,6 +31,7 @@ open_locale classical matrix polynomial noncomputable theory + open module.free polynomial matrix namespace linear_map diff --git a/src/linear_algebra/matrix/charpoly/coeff.lean b/src/linear_algebra/matrix/charpoly/coeff.lean index afe8e5c24be35..023aab0177aa1 100644 --- a/src/linear_algebra/matrix/charpoly/coeff.lean +++ b/src/linear_algebra/matrix/charpoly/coeff.lean @@ -4,15 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson, Jalex Stark -/ -import algebra.polynomial.big_operators -import data.matrix.char_p -import field_theory.finite.basic -import group_theory.perm.cycles +import data.polynomial.expand import linear_algebra.matrix.charpoly.basic -import linear_algebra.matrix.trace -import linear_algebra.matrix.to_lin -import ring_theory.polynomial.basic -import ring_theory.power_basis /-! # Characteristic polynomials @@ -182,54 +175,8 @@ begin not_false_iff] } end -@[simp] lemma finite_field.matrix.charpoly_pow_card {K : Type*} [field K] [fintype K] - (M : matrix n n K) : (M ^ (fintype.card K)).charpoly = M.charpoly := -begin - casesI (is_empty_or_nonempty n).symm, - { cases char_p.exists K with p hp, letI := hp, - rcases finite_field.card K p with ⟨⟨k, kpos⟩, ⟨hp, hk⟩⟩, - haveI : fact p.prime := ⟨hp⟩, - dsimp at hk, rw hk at *, - apply (frobenius_inj K[X] p).iterate k, - repeat { rw iterate_frobenius, rw ← hk }, - rw ← finite_field.expand_card, - unfold charpoly, rw [alg_hom.map_det, ← coe_det_monoid_hom, - ← (det_monoid_hom : matrix n n K[X] →* K[X]).map_pow], - apply congr_arg det, - refine mat_poly_equiv.injective _, - rw [alg_equiv.map_pow, mat_poly_equiv_charmatrix, hk, sub_pow_char_pow_of_commute, ← C_pow], - { exact (id (mat_poly_equiv_eq_X_pow_sub_C (p ^ k) M) : _) }, - { exact (C M).commute_X } }, - { -- TODO[gh-6025]: remove this `haveI` once `subsingleton_of_empty_right` is a global instance - haveI : subsingleton (matrix n n K) := subsingleton_of_empty_right, - exact congr_arg _ (subsingleton.elim _ _), }, -end - -@[simp] lemma zmod.charpoly_pow_card (M : matrix n n (zmod p)) : - (M ^ p).charpoly = M.charpoly := -by { have h := finite_field.matrix.charpoly_pow_card M, rwa zmod.card at h, } - -lemma finite_field.trace_pow_card {K : Type*} [field K] [fintype K] - (M : matrix n n K) : trace (M ^ (fintype.card K)) = trace M ^ (fintype.card K) := -begin - casesI is_empty_or_nonempty n, - { simp [zero_pow fintype.card_pos, matrix.trace], }, - rw [matrix.trace_eq_neg_charpoly_coeff, matrix.trace_eq_neg_charpoly_coeff, - finite_field.matrix.charpoly_pow_card, finite_field.pow_card] -end - -lemma zmod.trace_pow_card {p : ℕ} [fact p.prime] (M : matrix n n (zmod p)) : - trace (M ^ p) = (trace M)^p := -by { have h := finite_field.trace_pow_card M, rwa zmod.card at h, } - namespace matrix -theorem is_integral : is_integral R M := ⟨M.charpoly, ⟨charpoly_monic M, aeval_self_charpoly M⟩⟩ - -theorem minpoly_dvd_charpoly {K : Type*} [field K] (M : matrix n n K) : - (minpoly K M) ∣ M.charpoly := -minpoly.dvd _ _ (aeval_self_charpoly M) - /-- Any matrix polynomial `p` is equivalent under evaluation to `p %ₘ M.charpoly`; that is, `p` is equivalent to a polynomial with degree less than the dimension of the matrix. -/ lemma aeval_eq_aeval_mod_charpoly (M : matrix n n R) (p : R[X]) : @@ -243,29 +190,3 @@ lemma pow_eq_aeval_mod_charpoly (M : matrix n n R) (k : ℕ) : M^k = aeval M (X^ by rw [←aeval_eq_aeval_mod_charpoly, map_pow, aeval_X] end matrix - -section power_basis - -open algebra - -/-- The characteristic polynomial of the map `λ x, a * x` is the minimal polynomial of `a`. - -In combination with `det_eq_sign_charpoly_coeff` or `trace_eq_neg_charpoly_coeff` -and a bit of rewriting, this will allow us to conclude the -field norm resp. trace of `x` is the product resp. sum of `x`'s conjugates. --/ -lemma charpoly_left_mul_matrix {K S : Type*} [field K] [comm_ring S] [algebra K S] - (h : power_basis K S) : - (left_mul_matrix h.basis h.gen).charpoly = minpoly K h.gen := -begin - apply minpoly.unique, - { apply matrix.charpoly_monic }, - { apply (injective_iff_map_eq_zero (left_mul_matrix _)).mp (left_mul_matrix_injective h.basis), - rw [← polynomial.aeval_alg_hom_apply, aeval_self_charpoly] }, - { intros q q_monic root_q, - rw [matrix.charpoly_degree_eq_dim, fintype.card_fin, degree_eq_nat_degree q_monic.ne_zero], - apply with_bot.some_le_some.mpr, - exact h.dim_le_nat_degree_of_root q_monic.ne_zero root_q } -end - -end power_basis diff --git a/src/linear_algebra/matrix/charpoly/finite_field.lean b/src/linear_algebra/matrix/charpoly/finite_field.lean new file mode 100644 index 0000000000000..e701a3d07ed0d --- /dev/null +++ b/src/linear_algebra/matrix/charpoly/finite_field.lean @@ -0,0 +1,60 @@ +/- +Copyright (c) 2020 Aaron Anderson, Jalex Stark. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Aaron Anderson, Jalex Stark +-/ +import linear_algebra.matrix.charpoly.coeff +import field_theory.finite.basic +import data.matrix.char_p + + +/-! +# Results on characteristic polynomials and traces over finite fields. +-/ + +noncomputable theory + +open polynomial matrix +open_locale polynomial + +variables {n : Type*} [decidable_eq n] [fintype n] + +@[simp] lemma finite_field.matrix.charpoly_pow_card {K : Type*} [field K] [fintype K] + (M : matrix n n K) : (M ^ (fintype.card K)).charpoly = M.charpoly := +begin + casesI (is_empty_or_nonempty n).symm, + { cases char_p.exists K with p hp, letI := hp, + rcases finite_field.card K p with ⟨⟨k, kpos⟩, ⟨hp, hk⟩⟩, + haveI : fact p.prime := ⟨hp⟩, + dsimp at hk, rw hk at *, + apply (frobenius_inj K[X] p).iterate k, + repeat { rw iterate_frobenius, rw ← hk }, + rw ← finite_field.expand_card, + unfold charpoly, rw [alg_hom.map_det, ← coe_det_monoid_hom, + ← (det_monoid_hom : matrix n n K[X] →* K[X]).map_pow], + apply congr_arg det, + refine mat_poly_equiv.injective _, + rw [alg_equiv.map_pow, mat_poly_equiv_charmatrix, hk, sub_pow_char_pow_of_commute, ← C_pow], + { exact (id (mat_poly_equiv_eq_X_pow_sub_C (p ^ k) M) : _) }, + { exact (C M).commute_X } }, + { -- TODO[gh-6025]: remove this `haveI` once `subsingleton_of_empty_right` is a global instance + haveI : subsingleton (matrix n n K) := subsingleton_of_empty_right, + exact congr_arg _ (subsingleton.elim _ _), }, +end + +@[simp] lemma zmod.charpoly_pow_card {p : ℕ} [fact p.prime] (M : matrix n n (zmod p)) : + (M ^ p).charpoly = M.charpoly := +by { have h := finite_field.matrix.charpoly_pow_card M, rwa zmod.card at h, } + +lemma finite_field.trace_pow_card {K : Type*} [field K] [fintype K] + (M : matrix n n K) : trace (M ^ (fintype.card K)) = trace M ^ (fintype.card K) := +begin + casesI is_empty_or_nonempty n, + { simp [zero_pow fintype.card_pos, matrix.trace], }, + rw [matrix.trace_eq_neg_charpoly_coeff, matrix.trace_eq_neg_charpoly_coeff, + finite_field.matrix.charpoly_pow_card, finite_field.pow_card] +end + +lemma zmod.trace_pow_card {p : ℕ} [fact p.prime] (M : matrix n n (zmod p)) : + trace (M ^ p) = (trace M)^p := +by { have h := finite_field.trace_pow_card M, rwa zmod.card at h, } diff --git a/src/linear_algebra/matrix/charpoly/minpoly.lean b/src/linear_algebra/matrix/charpoly/minpoly.lean new file mode 100644 index 0000000000000..d65cfc36474c3 --- /dev/null +++ b/src/linear_algebra/matrix/charpoly/minpoly.lean @@ -0,0 +1,61 @@ +/- +Copyright (c) 2020 Aaron Anderson, Jalex Stark. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Aaron Anderson, Jalex Stark +-/ + +import linear_algebra.matrix.charpoly.coeff +import ring_theory.power_basis + +/-! +# The minimal polynomial divides the characteristic polynomial of a matrix. +-/ + +noncomputable theory + +universes u v + +open polynomial matrix + +variables {R : Type u} [comm_ring R] +variables {n : Type v} [decidable_eq n] [fintype n] + +open finset + +variable {M : matrix n n R} + +namespace matrix + +theorem is_integral : is_integral R M := ⟨M.charpoly, ⟨charpoly_monic M, aeval_self_charpoly M⟩⟩ + +theorem minpoly_dvd_charpoly {K : Type*} [field K] (M : matrix n n K) : + (minpoly K M) ∣ M.charpoly := +minpoly.dvd _ _ (aeval_self_charpoly M) + +end matrix + +section power_basis + +open algebra + +/-- The characteristic polynomial of the map `λ x, a * x` is the minimal polynomial of `a`. + +In combination with `det_eq_sign_charpoly_coeff` or `trace_eq_neg_charpoly_coeff` +and a bit of rewriting, this will allow us to conclude the +field norm resp. trace of `x` is the product resp. sum of `x`'s conjugates. +-/ +lemma charpoly_left_mul_matrix {K S : Type*} [field K] [comm_ring S] [algebra K S] + (h : power_basis K S) : + (left_mul_matrix h.basis h.gen).charpoly = minpoly K h.gen := +begin + apply minpoly.unique, + { apply matrix.charpoly_monic }, + { apply (injective_iff_map_eq_zero (left_mul_matrix _)).mp (left_mul_matrix_injective h.basis), + rw [← polynomial.aeval_alg_hom_apply, aeval_self_charpoly] }, + { intros q q_monic root_q, + rw [matrix.charpoly_degree_eq_dim, fintype.card_fin, degree_eq_nat_degree q_monic.ne_zero], + apply with_bot.some_le_some.mpr, + exact h.dim_le_nat_degree_of_root q_monic.ne_zero root_q } +end + +end power_basis diff --git a/src/ring_theory/norm.lean b/src/ring_theory/norm.lean index 044e4098ad177..1ccbfdd353385 100644 --- a/src/ring_theory/norm.lean +++ b/src/ring_theory/norm.lean @@ -6,7 +6,7 @@ Authors: Anne Baanen import field_theory.primitive_element import linear_algebra.determinant -import linear_algebra.matrix.charpoly.coeff +import linear_algebra.matrix.charpoly.minpoly import linear_algebra.matrix.to_linear_equiv import field_theory.is_alg_closed.algebraic_closure diff --git a/src/ring_theory/trace.lean b/src/ring_theory/trace.lean index 69651a2c3beb4..a00ca29fbb241 100644 --- a/src/ring_theory/trace.lean +++ b/src/ring_theory/trace.lean @@ -5,7 +5,7 @@ Authors: Anne Baanen -/ import linear_algebra.matrix.bilinear_form -import linear_algebra.matrix.charpoly.coeff +import linear_algebra.matrix.charpoly.minpoly import linear_algebra.determinant import linear_algebra.vandermonde import linear_algebra.trace From 31529827d0f68d1fbd429edc393a928f677f4aba Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 4 May 2022 17:58:28 +0000 Subject: [PATCH 111/153] feat(representation/Rep): linear structures (#13782) Make `Rep k G` a `k`-linear (and `k`-linear monoidal) category. Co-authored-by: Scott Morrison --- src/algebra/category/Module/monoidal.lean | 6 +- src/category_theory/linear/default.lean | 2 + .../linear/functor_category.lean | 54 +++++++++ src/category_theory/monoidal/linear.lean | 45 ++++++++ src/category_theory/monoidal/transport.lean | 11 +- .../preadditive/functor_category.lean | 2 +- src/representation_theory/Action.lean | 107 +++++++++++++++++- src/representation_theory/Rep.lean | 10 +- 8 files changed, 228 insertions(+), 9 deletions(-) create mode 100644 src/category_theory/linear/functor_category.lean create mode 100644 src/category_theory/monoidal/linear.lean diff --git a/src/algebra/category/Module/monoidal.lean b/src/algebra/category/Module/monoidal.lean index cfdf034c416d1..6f8a93b06d79f 100644 --- a/src/algebra/category/Module/monoidal.lean +++ b/src/algebra/category/Module/monoidal.lean @@ -8,7 +8,7 @@ import category_theory.closed.monoidal import algebra.category.Module.basic import linear_algebra.tensor_product import category_theory.linear.yoneda -import category_theory.monoidal.preadditive +import category_theory.monoidal.linear /-! # The symmetric monoidal category structure on R-modules @@ -273,6 +273,10 @@ instance : monoidal_preadditive (Module.{u} R) := tensor_add' := by { intros, ext, simp [tensor_product.tmul_add], }, add_tensor' := by { intros, ext, simp [tensor_product.add_tmul], }, } +instance : monoidal_linear R (Module.{u} R) := +{ tensor_smul' := by { intros, ext, simp, }, + smul_tensor' := by { intros, ext, simp [tensor_product.smul_tmul], }, } + /-- Auxiliary definition for the `monoidal_closed` instance on `Module R`. (This is only a separate definition in order to speed up typechecking. ) diff --git a/src/category_theory/linear/default.lean b/src/category_theory/linear/default.lean index 49491eac09963..ea1e87f9ee9dd 100644 --- a/src/category_theory/linear/default.lean +++ b/src/category_theory/linear/default.lean @@ -15,6 +15,8 @@ import algebra.algebra.basic An `R`-linear category is a category in which `X ⟶ Y` is an `R`-module in such a way that composition of morphisms is `R`-linear in both variables. +Note that sometimes in the literature a "linear category" is further required to be abelian. + ## Implementation Corresponding to the fact that we need to have an `add_comm_group X` structure in place diff --git a/src/category_theory/linear/functor_category.lean b/src/category_theory/linear/functor_category.lean new file mode 100644 index 0000000000000..82e68a1a6f28b --- /dev/null +++ b/src/category_theory/linear/functor_category.lean @@ -0,0 +1,54 @@ +/- +Copyright (c) 2022 Scott Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Scott Morrison +-/ +import category_theory.preadditive.functor_category +import category_theory.linear.default + +/-! +# Linear structure on functor categories + +If `C` and `D` are categories and `D` is `R`-linear, +then `C ⥤ D` is also `R`-linear. + +-/ + +open_locale big_operators + +namespace category_theory +open category_theory.limits linear + +variables {R : Type*} [semiring R] +variables {C D : Type*} [category C] [category D] [preadditive D] [linear R D] + +instance functor_category_linear : linear R (C ⥤ D) := +{ hom_module := λ F G, + { smul := λ r α, + { app := λ X, r • α.app X, + naturality' := by { intros, rw [comp_smul, smul_comp, α.naturality] } }, + one_smul := by { intros, ext, apply one_smul }, + zero_smul := by { intros, ext, apply zero_smul }, + smul_zero := by { intros, ext, apply smul_zero }, + add_smul := by { intros, ext, apply add_smul }, + smul_add := by { intros, ext, apply smul_add }, + mul_smul := by { intros, ext, apply mul_smul } }, + smul_comp' := by { intros, ext, apply smul_comp }, + comp_smul' := by { intros, ext, apply comp_smul } } + +namespace nat_trans + +variables {F G : C ⥤ D} + +/-- Application of a natural transformation at a fixed object, +as group homomorphism -/ +@[simps] def app_linear_map (X : C) : (F ⟶ G) →ₗ[R] (F.obj X ⟶ G.obj X) := +{ to_fun := λ α, α.app X, + map_add' := λ _ _, rfl, + map_smul' := λ _ _, rfl, } + +@[simp] lemma app_smul (X : C) (r : R) (α : F ⟶ G) : (r • α).app X = r • α.app X := rfl + +end nat_trans + +end category_theory diff --git a/src/category_theory/monoidal/linear.lean b/src/category_theory/monoidal/linear.lean new file mode 100644 index 0000000000000..30269ad6723b5 --- /dev/null +++ b/src/category_theory/monoidal/linear.lean @@ -0,0 +1,45 @@ +/- +Copyright (c) 2022 Scott Morrison. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Scott Morrison +-/ +import category_theory.linear.linear_functor +import category_theory.monoidal.preadditive + +/-! +# Linear monoidal categories + +A monoidal category is `monoidal_linear R` if it is monoidal preadditive and +tensor product of morphisms is `R`-linear in both factors. +-/ + +namespace category_theory + +open category_theory.limits +open category_theory.monoidal_category + +variables (R : Type*) [semiring R] +variables (C : Type*) [category C] [preadditive C] [linear R C] +variables [monoidal_category C] [monoidal_preadditive C] + +/-- +A category is `monoidal_linear R` if tensoring is `R`-linear in both factors. +-/ +class monoidal_linear := +(tensor_smul' : ∀ {W X Y Z : C} (f : W ⟶ X) (r : R) (g : Y ⟶ Z), + f ⊗ (r • g) = r • (f ⊗ g) . obviously) +(smul_tensor' : ∀ {W X Y Z : C} (r : R) (f : W ⟶ X) (g : Y ⟶ Z), + (r • f) ⊗ g = r • (f ⊗ g) . obviously) + +restate_axiom monoidal_linear.tensor_smul' +restate_axiom monoidal_linear.smul_tensor' +attribute [simp] monoidal_linear.tensor_smul monoidal_linear.smul_tensor + +variables [monoidal_linear R C] + +instance tensor_left_linear (X : C) : (tensor_left X).linear R := {} +instance tensor_right_linear (X : C) : (tensor_right X).linear R := {} +instance tensoring_left_linear (X : C) : ((tensoring_left C).obj X).linear R := {} +instance tensoring_right_linear (X : C) : ((tensoring_right C).obj X).linear R := {} + +end category_theory diff --git a/src/category_theory/monoidal/transport.lean b/src/category_theory/monoidal/transport.lean index 0f56702cb230f..954619dd50498 100644 --- a/src/category_theory/monoidal/transport.lean +++ b/src/category_theory/monoidal/transport.lean @@ -32,7 +32,7 @@ variables {D : Type u₂} [category.{v₂} D] /-- Transport a monoidal structure along an equivalence of (plain) categories. -/ -@[simps] +@[simps {attrs := [`_refl_lemma]}] -- We just want these simp lemmas locally def transport (e : C ≌ D) : monoidal_category.{v₂} D := { tensor_obj := λ X Y, e.functor.obj (e.inverse.obj X ⊗ e.inverse.obj Y), tensor_hom := λ W X Y Z f g, e.functor.map (e.inverse.map f ⊗ e.inverse.map g), @@ -137,6 +137,13 @@ def transported (e : C ≌ D) := D instance (e : C ≌ D) : monoidal_category (transported e) := transport e instance (e : C ≌ D) : inhabited (transported e) := ⟨𝟙_ _⟩ +section +local attribute [simp] transport_tensor_unit + +section +local attribute [simp] transport_tensor_hom transport_associator + transport_left_unitor transport_right_unitor + /-- We can upgrade `e.functor` to a lax monoidal functor from `C` to `D` with the transported structure. -/ @@ -191,6 +198,7 @@ def lax_to_transported (e : C ≌ D) : lax_monoidal_functor C (transported e) := congr' 1, simp only [←right_unitor_naturality, id_comp, ←tensor_comp_assoc, comp_id], end, }. +end /-- We can upgrade `e.functor` to a monoidal functor from `C` to `D` with the transported structure. @@ -200,6 +208,7 @@ def to_transported (e : C ≌ D) : monoidal_functor C (transported e) := { to_lax_monoidal_functor := lax_to_transported e, ε_is_iso := by { dsimp, apply_instance, }, μ_is_iso := λ X Y, by { dsimp, apply_instance, }, } +end instance (e : C ≌ D) : is_equivalence (to_transported e).to_functor := by { dsimp, apply_instance, } diff --git a/src/category_theory/preadditive/functor_category.lean b/src/category_theory/preadditive/functor_category.lean index e6c88afc29997..1d6dc357ec7db 100644 --- a/src/category_theory/preadditive/functor_category.lean +++ b/src/category_theory/preadditive/functor_category.lean @@ -21,7 +21,7 @@ open category_theory.limits preadditive variables {C D : Type*} [category C] [category D] [preadditive D] -instance : preadditive (C ⥤ D) := +instance functor_category_preadditive : preadditive (C ⥤ D) := { hom_group := λ F G, { add := λ α β, { app := λ X, α.app X + β.app X, diff --git a/src/representation_theory/Action.lean b/src/representation_theory/Action.lean index 6e39809bb4fff..d5b408ec7cb18 100644 --- a/src/representation_theory/Action.lean +++ b/src/representation_theory/Action.lean @@ -10,9 +10,11 @@ import category_theory.limits.preserves.basic import category_theory.adjunction.limits import category_theory.monoidal.functor_category import category_theory.monoidal.transport +import category_theory.monoidal.linear import category_theory.monoidal.braided import category_theory.abelian.functor_category import category_theory.abelian.transfer +import category_theory.linear.functor_category /-! # `Action V G`, the category of actions of a monoid `G` inside some category `V`. @@ -25,7 +27,7 @@ and construct the restriction functors `res {G H : Mon} (f : G ⟶ H) : Action V * When `V` has (co)limits so does `Action V G`. * When `V` is monoidal, braided, or symmetric, so is `Action V G`. -* When `V` is preadditive or abelian so is `Action V G`. +* When `V` is preadditive, linear, or abelian so is `Action V G`. -/ universes u @@ -272,8 +274,33 @@ instance : preadditive (Action V G) := instance : functor.additive (functor_category_equivalence V G).functor := {} +@[simp] lemma zero_hom {X Y : Action V G} : (0 : X ⟶ Y).hom = 0 := rfl +@[simp] lemma neg_hom {X Y : Action V G} (f : X ⟶ Y) : (-f).hom = -f.hom := rfl +@[simp] lemma add_hom {X Y : Action V G} (f g : X ⟶ Y) : (f + g).hom = f.hom + g.hom := rfl + end preadditive +section linear +variables [preadditive V] {R : Type*} [semiring R] [linear R V] + +instance : linear R (Action V G) := +{ hom_module := λ X Y, + { smul := λ r f, ⟨r • f.hom, by simp [f.comm]⟩, + one_smul := by { intros, ext, exact one_smul _ _, }, + smul_zero := by { intros, ext, exact smul_zero _, }, + zero_smul := by { intros, ext, exact zero_smul _ _, }, + add_smul := by { intros, ext, exact add_smul _ _ _, }, + smul_add := by { intros, ext, exact smul_add _ _ _, }, + mul_smul := by { intros, ext, exact mul_smul _ _ _, }, }, + smul_comp' := by { intros, ext, exact linear.smul_comp _ _ _ _ _ _, }, + comp_smul' := by { intros, ext, exact linear.comp_smul _ _ _ _ _ _, }, } + +instance : functor.linear R (functor_category_equivalence V G).functor := {} + +@[simp] lemma smul_hom {X Y : Action V G} (r : R) (f : X ⟶ Y) : (r • f).hom = r • f.hom := rfl + +end linear + section abelian /-- Auxilliary construction for the `abelian (Action V G)` instance. -/ def abelian_aux : Action V G ≌ (ulift.{u} (single_obj G) ⥤ V) := @@ -290,6 +317,47 @@ variables [monoidal_category V] instance : monoidal_category (Action V G) := monoidal.transport (Action.functor_category_equivalence _ _).symm +@[simp] lemma tensor_V {X Y : Action V G} : (X ⊗ Y).V = X.V ⊗ Y.V := rfl +@[simp] lemma tensor_rho {X Y : Action V G} {g : G} : (X ⊗ Y).ρ g = X.ρ g ⊗ Y.ρ g := rfl +@[simp] lemma tensor_hom {W X Y Z : Action V G} (f : W ⟶ X) (g : Y ⟶ Z) : + (f ⊗ g).hom = f.hom ⊗ g.hom := rfl +@[simp] lemma associator_hom_hom {X Y Z : Action V G} : + hom.hom (α_ X Y Z).hom = (α_ X.V Y.V Z.V).hom := +begin + dsimp [monoidal.transport_associator], + simp, +end +@[simp] lemma associator_inv_hom {X Y Z : Action V G} : + hom.hom (α_ X Y Z).inv = (α_ X.V Y.V Z.V).inv := +begin + dsimp [monoidal.transport_associator], + simp, +end +@[simp] lemma left_unitor_hom_hom {X : Action V G} : + hom.hom (λ_ X).hom = (λ_ X.V).hom := +begin + dsimp [monoidal.transport_left_unitor], + simp, +end +@[simp] lemma left_unitor_inv_hom {X : Action V G} : + hom.hom (λ_ X).inv = (λ_ X.V).inv := +begin + dsimp [monoidal.transport_left_unitor], + simp, +end +@[simp] lemma right_unitor_hom_hom {X : Action V G} : + hom.hom (ρ_ X).hom = (ρ_ X.V).hom := +begin + dsimp [monoidal.transport_right_unitor], + simp, +end +@[simp] lemma right_unitor_inv_hom {X : Action V G} : + hom.hom (ρ_ X).inv = (ρ_ X.V).inv := +begin + dsimp [monoidal.transport_right_unitor], + simp, +end + variables (V G) /-- When `V` is monoidal the forgetful functor `Action V G` to `V` is monoidal. -/ @@ -302,20 +370,35 @@ def forget_monoidal : monoidal_functor (Action V G) V := instance forget_monoidal_faithful : faithful (forget_monoidal V G).to_functor := by { change faithful (forget V G), apply_instance, } -instance [braided_category V] : braided_category (Action V G) := +section +variables [braided_category V] + +instance : braided_category (Action V G) := braided_category_of_faithful (forget_monoidal V G) (λ X Y, mk_iso (β_ _ _) (by tidy)) (by tidy) /-- When `V` is braided the forgetful functor `Action V G` to `V` is braided. -/ @[simps] -def forget_braided [braided_category V] : braided_functor (Action V G) V := +def forget_braided : braided_functor (Action V G) V := { ..forget_monoidal _ _, } -instance forget_braided_faithful [braided_category V] : faithful (forget_braided V G).to_functor := +instance forget_braided_faithful : faithful (forget_braided V G).to_functor := by { change faithful (forget V G), apply_instance, } +end + instance [symmetric_category V] : symmetric_category (Action V G) := symmetric_category_of_faithful (forget_braided V G) +variables [preadditive V] [monoidal_preadditive V] + +local attribute [simp] monoidal_preadditive.tensor_add monoidal_preadditive.add_tensor + +instance : monoidal_preadditive (Action V G) := {} + +variables {R : Type*} [semiring R] [linear R V] [monoidal_linear R V] + +instance : monoidal_linear R (Action V G) := {} + end monoidal /-- Actions/representations of the trivial group are just objects in the ambient category. -/ @@ -365,6 +448,14 @@ attribute [simps] res_comp -- TODO promote `res` to a pseudofunctor from -- the locally discrete bicategory constructed from `Monᵒᵖ` to `Cat`, sending `G` to `Action V G`. +variables {G} {H : Mon.{u}} (f : G ⟶ H) + +instance res_additive [preadditive V] : (res V f).additive := {} + +variables {R : Type*} [semiring R] + +instance res_linear [preadditive V] [linear R V] : (res V f).linear R := {} + end Action namespace category_theory.functor @@ -387,4 +478,12 @@ def map_Action (F : V ⥤ W) (G : Mon.{u}) : Action V G ⥤ Action W G := map_id' := λ M, by { ext, simp only [Action.id_hom, F.map_id], }, map_comp' := λ M N P f g, by { ext, simp only [Action.comp_hom, F.map_comp], }, } +variables (F : V ⥤ W) (G : Mon.{u}) [preadditive V] [preadditive W] + +instance map_Action_preadditive [F.additive] : (F.map_Action G).additive := {} + +variables {R : Type*} [semiring R] [category_theory.linear R V] [category_theory.linear R W] + +instance map_Action_linear [F.additive] [F.linear R] : (F.map_Action G).linear R := {} + end category_theory.functor diff --git a/src/representation_theory/Rep.lean b/src/representation_theory/Rep.lean index c9c46702f7114..ad9e093c6632d 100644 --- a/src/representation_theory/Rep.lean +++ b/src/representation_theory/Rep.lean @@ -18,7 +18,7 @@ Also `V.ρ` gives the homomorphism `G →* (V →ₗ[k] V)`. Conversely, given a homomorphism `ρ : G →* (V →ₗ[k] V)`, you can construct the bundled representation as `Rep.of ρ`. -We verify that `Rep k G` is an abelian symmetric monoidal category with all (co)limits. +We verify that `Rep k G` is a `k`-linear abelian symmetric monoidal category with all (co)limits. -/ universes u @@ -27,10 +27,14 @@ open category_theory open category_theory.limits /-- The category of `k`-linear representations of a monoid `G`. -/ -@[derive [large_category, concrete_category, has_limits, has_colimits, abelian]] +@[derive [large_category, concrete_category, has_limits, has_colimits, + preadditive, abelian]] abbreviation Rep (k G : Type u) [ring k] [monoid G] := Action (Module.{u} k) (Mon.of G) +instance (k G : Type u) [comm_ring k] [monoid G] : linear k (Rep k G) := +by apply_instance + namespace Rep variables {k G : Type u} [ring k] [monoid G] @@ -64,5 +68,7 @@ variables {k G : Type u} [comm_ring k] [monoid G] -- Verify that the symmetric monoidal structure is available. example : symmetric_category (Rep k G) := by apply_instance +example : monoidal_preadditive (Rep k G) := by apply_instance +example : monoidal_linear k (Rep k G) := by apply_instance end Rep From 209bb5d4dcf7f692be8bf6b57ccbc81b5eb35512 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 4 May 2022 17:58:29 +0000 Subject: [PATCH 112/153] feat(set_theory/game/{pgame, basic}): Add more order lemmas (#13807) --- src/set_theory/game/basic.lean | 6 ++++++ src/set_theory/game/pgame.lean | 13 +++++++++++++ 2 files changed, 19 insertions(+) diff --git a/src/set_theory/game/basic.lean b/src/set_theory/game/basic.lean index 259245dabdb20..254cc2f383344 100644 --- a/src/set_theory/game/basic.lean +++ b/src/set_theory/game/basic.lean @@ -67,6 +67,12 @@ into a `def` instead. -/ instance : has_lt game := ⟨quotient.lift₂ (λ x y, x < y) (λ x₁ y₁ x₂ y₂ hx hy, propext (lt_congr hx hy))⟩ +theorem lt_or_eq_of_le : ∀ {x y : game}, x ≤ y → x < y ∨ x = y := +by { rintro ⟨x⟩ ⟨y⟩, change _ → _ ∨ ⟦x⟧ = ⟦y⟧, rw quotient.eq, exact lt_or_equiv_of_le } + +instance : is_trichotomous game (<) := +⟨by { rintro ⟨x⟩ ⟨y⟩, change _ ∨ ⟦x⟧ = ⟦y⟧ ∨ _, rw quotient.eq, apply lt_or_equiv_or_gt }⟩ + @[simp] theorem not_le : ∀ {x y : game}, ¬ x ≤ y ↔ y < x := by { rintro ⟨x⟩ ⟨y⟩, exact not_le } diff --git a/src/set_theory/game/pgame.lean b/src/set_theory/game/pgame.lean index 7fbfbb9587a4b..c119b26854ca3 100644 --- a/src/set_theory/game/pgame.lean +++ b/src/set_theory/game/pgame.lean @@ -526,6 +526,19 @@ theorem le_congr {x₁ y₁ x₂ y₂} : x₁ ≈ x₂ → y₁ ≈ y₂ → (x theorem lt_congr {x₁ y₁ x₂ y₂} (hx : x₁ ≈ x₂) (hy : y₁ ≈ y₂) : x₁ < y₁ ↔ x₂ < y₂ := not_le.symm.trans $ (not_congr (le_congr hy hx)).trans not_le +theorem lt_or_equiv_of_le {x y : pgame} (h : x ≤ y) : x < y ∨ x ≈ y := +or_iff_not_imp_left.2 $ λ h', ⟨h, not_lt.1 h'⟩ + +theorem lt_or_equiv_or_gt (x y : pgame) : x < y ∨ x ≈ y ∨ y < x := +begin + by_cases h : x < y, + { exact or.inl h }, + { right, + cases (lt_or_equiv_of_le (not_lt.1 h)) with h' h', + { exact or.inr h' }, + { exact or.inl h'.symm } } +end + theorem equiv_congr_left {y₁ y₂} : y₁ ≈ y₂ ↔ ∀ x₁, x₁ ≈ y₁ ↔ x₁ ≈ y₂ := ⟨λ h x₁, ⟨λ h', equiv_trans h' h, λ h', equiv_trans h' (equiv_symm h)⟩, λ h, (h y₁).1 $ equiv_refl _⟩ From aabcd8950d0738c0f5ae431461a488ac439c70bc Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Wed, 4 May 2022 17:58:30 +0000 Subject: [PATCH 113/153] chore(analysis/analytic_composition): weaken some typeclass arguments (#13924) There's no need to do a long computation to show the multilinear_map is bounded, when continuity follows directly from the definition. This deletes `comp_along_composition_aux`, and moves the lemmas about the norm of `comp_along_composition` further down the file so as to get the lemmas with weaker typeclass requirements out of the way first. The norm proofs are essentially unchanged. --- src/analysis/analytic/composition.lean | 142 +++++++++--------- .../normed_space/indicator_function.lean | 2 +- 2 files changed, 75 insertions(+), 69 deletions(-) diff --git a/src/analysis/analytic/composition.lean b/src/analysis/analytic/composition.lean index 73a7fce074e95..e7d04555e7627 100644 --- a/src/analysis/analytic/composition.lean +++ b/src/analysis/analytic/composition.lean @@ -65,18 +65,21 @@ in more details below in the paragraph on associativity. noncomputable theory -variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] -{E : Type*} [normed_group E] [normed_space 𝕜 E] -{F : Type*} [normed_group F] [normed_space 𝕜 F] -{G : Type*} [normed_group G] [normed_space 𝕜 G] -{H : Type*} [normed_group H] [normed_space 𝕜 H] - +variables {𝕜 : Type*} {E F G H : Type*} open filter list open_locale topological_space big_operators classical nnreal ennreal +section topological +variables [comm_ring 𝕜] [add_comm_group E] [add_comm_group F] [add_comm_group G] +variables [module 𝕜 E] [module 𝕜 F] [module 𝕜 G] +variables [topological_space E] [topological_space F] [topological_space G] + /-! ### Composing formal multilinear series -/ namespace formal_multilinear_series +variables [topological_add_group E] [has_continuous_const_smul 𝕜 E] +variables [topological_add_group F] [has_continuous_const_smul 𝕜 F] +variables [topological_add_group G] [has_continuous_const_smul 𝕜 G] /-! In this paragraph, we define the composition of formal multilinear series, by summing over all @@ -166,58 +169,24 @@ end formal_multilinear_series namespace continuous_multilinear_map open formal_multilinear_series - -/-- Given a formal multilinear series `p`, a composition `c` of `n` and a continuous multilinear -map `f` in `c.length` variables, one may form a multilinear map in `n` variables by applying -the right coefficient of `p` to each block of the composition, and then applying `f` to the -resulting vector. It is called `f.comp_along_composition_aux p c`. -This function admits a version as a continuous multilinear map, called -`f.comp_along_composition p c` below. -/ -def comp_along_composition_aux {n : ℕ} - (p : formal_multilinear_series 𝕜 E F) (c : composition n) - (f : continuous_multilinear_map 𝕜 (λ (i : fin c.length), F) G) : - multilinear_map 𝕜 (λ i : fin n, E) G := -{ to_fun := λ v, f (p.apply_composition c v), - map_add' := λ v i x y, by simp only [apply_composition_update, - continuous_multilinear_map.map_add], - map_smul' := λ v i c x, by simp only [apply_composition_update, - continuous_multilinear_map.map_smul] } - -/-- The norm of `f.comp_along_composition_aux p c` is controlled by the product of -the norms of the relevant bits of `f` and `p`. -/ -lemma comp_along_composition_aux_bound {n : ℕ} - (p : formal_multilinear_series 𝕜 E F) (c : composition n) - (f : continuous_multilinear_map 𝕜 (λ (i : fin c.length), F) G) (v : fin n → E) : - ∥f.comp_along_composition_aux p c v∥ ≤ - ∥f∥ * (∏ i, ∥p (c.blocks_fun i)∥) * (∏ i : fin n, ∥v i∥) := -calc ∥f.comp_along_composition_aux p c v∥ = ∥f (p.apply_composition c v)∥ : rfl -... ≤ ∥f∥ * ∏ i, ∥p.apply_composition c v i∥ : continuous_multilinear_map.le_op_norm _ _ -... ≤ ∥f∥ * ∏ i, ∥p (c.blocks_fun i)∥ * - ∏ j : fin (c.blocks_fun i), ∥(v ∘ (c.embedding i)) j∥ : - begin - apply mul_le_mul_of_nonneg_left _ (norm_nonneg _), - refine finset.prod_le_prod (λ i hi, norm_nonneg _) (λ i hi, _), - apply continuous_multilinear_map.le_op_norm, - end -... = ∥f∥ * (∏ i, ∥p (c.blocks_fun i)∥) * - ∏ i (j : fin (c.blocks_fun i)), ∥(v ∘ (c.embedding i)) j∥ : - by rw [finset.prod_mul_distrib, mul_assoc] -... = ∥f∥ * (∏ i, ∥p (c.blocks_fun i)∥) * (∏ i : fin n, ∥v i∥) : - by { rw [← c.blocks_fin_equiv.prod_comp, ← finset.univ_sigma_univ, finset.prod_sigma], - congr } +variables [topological_add_group E] [has_continuous_const_smul 𝕜 E] +variables [topological_add_group F] [has_continuous_const_smul 𝕜 F] /-- Given a formal multilinear series `p`, a composition `c` of `n` and a continuous multilinear map `f` in `c.length` variables, one may form a continuous multilinear map in `n` variables by applying the right coefficient of `p` to each block of the composition, and then applying `f` to -the resulting vector. It is called `f.comp_along_composition p c`. It is constructed from the -analogous multilinear function `f.comp_along_composition_aux p c`, together with a norm -control to get the continuity. -/ +the resulting vector. It is called `f.comp_along_composition p c`. -/ def comp_along_composition {n : ℕ} (p : formal_multilinear_series 𝕜 E F) (c : composition n) (f : continuous_multilinear_map 𝕜 (λ (i : fin c.length), F) G) : continuous_multilinear_map 𝕜 (λ i : fin n, E) G := -(f.comp_along_composition_aux p c).mk_continuous _ - (f.comp_along_composition_aux_bound p c) +{ to_fun := λ v, f (p.apply_composition c v), + map_add' := λ v i x y, by simp only [apply_composition_update, + continuous_multilinear_map.map_add], + map_smul' := λ v i c x, by simp only [apply_composition_update, + continuous_multilinear_map.map_smul], + cont := f.cont.comp $ continuous_pi $ λ i, (coe_continuous _).comp $ continuous_pi $ λ j, + continuous_apply _, } @[simp] lemma comp_along_composition_apply {n : ℕ} (p : formal_multilinear_series 𝕜 E F) (c : composition n) @@ -227,13 +196,14 @@ def comp_along_composition {n : ℕ} end continuous_multilinear_map namespace formal_multilinear_series +variables [topological_add_group E] [has_continuous_const_smul 𝕜 E] +variables [topological_add_group F] [has_continuous_const_smul 𝕜 F] +variables [topological_add_group G] [has_continuous_const_smul 𝕜 G] /-- Given two formal multilinear series `q` and `p` and a composition `c` of `n`, one may form a continuous multilinear map in `n` variables by applying the right coefficient of `p` to each block of the composition, and then applying `q c.length` to the resulting vector. It is -called `q.comp_along_composition p c`. It is constructed from the analogous multilinear -function `q.comp_along_composition_aux p c`, together with a norm control to get -the continuity. -/ +called `q.comp_along_composition p c`. -/ def comp_along_composition {n : ℕ} (q : formal_multilinear_series 𝕜 F G) (p : formal_multilinear_series 𝕜 E F) (c : composition n) : continuous_multilinear_map 𝕜 (λ i : fin n, E) G := @@ -244,21 +214,6 @@ def comp_along_composition {n : ℕ} (c : composition n) (v : fin n → E) : (q.comp_along_composition p c) v = q c.length (p.apply_composition c v) := rfl -/-- The norm of `q.comp_along_composition p c` is controlled by the product of -the norms of the relevant bits of `q` and `p`. -/ -lemma comp_along_composition_norm {n : ℕ} - (q : formal_multilinear_series 𝕜 F G) (p : formal_multilinear_series 𝕜 E F) - (c : composition n) : - ∥q.comp_along_composition p c∥ ≤ ∥q c.length∥ * ∏ i, ∥p (c.blocks_fun i)∥ := -multilinear_map.mk_continuous_norm_le _ - (mul_nonneg (norm_nonneg _) (finset.prod_nonneg (λ i hi, norm_nonneg _))) _ - -lemma comp_along_composition_nnnorm {n : ℕ} - (q : formal_multilinear_series 𝕜 F G) (p : formal_multilinear_series 𝕜 E F) - (c : composition n) : - ∥q.comp_along_composition p c∥₊ ≤ ∥q c.length∥₊ * ∏ i, ∥p (c.blocks_fun i)∥₊ := -by { rw ← nnreal.coe_le_coe, push_cast, exact q.comp_along_composition_norm p c } - /-- Formal composition of two formal multilinear series. The `n`-th coefficient in the composition is defined to be the sum of `q.comp_along_composition p c` over all compositions of `n`. In other words, this term (as a multilinear function applied to `v_0, ..., v_{n-1}`) is @@ -329,6 +284,57 @@ end q.comp p.remove_zero = q.comp p := by { ext n, simp [formal_multilinear_series.comp] } +end formal_multilinear_series + +end topological + +variables [nondiscrete_normed_field 𝕜] + [normed_group E] [normed_space 𝕜 E] + [normed_group F] [normed_space 𝕜 F] + [normed_group G] [normed_space 𝕜 G] + [normed_group H] [normed_space 𝕜 H] + +namespace formal_multilinear_series + +/-- The norm of `f.comp_along_composition p c` is controlled by the product of +the norms of the relevant bits of `f` and `p`. -/ +lemma comp_along_composition_bound {n : ℕ} + (p : formal_multilinear_series 𝕜 E F) (c : composition n) + (f : continuous_multilinear_map 𝕜 (λ (i : fin c.length), F) G) (v : fin n → E) : + ∥f.comp_along_composition p c v∥ ≤ + ∥f∥ * (∏ i, ∥p (c.blocks_fun i)∥) * (∏ i : fin n, ∥v i∥) := +calc ∥f.comp_along_composition p c v∥ = ∥f (p.apply_composition c v)∥ : rfl +... ≤ ∥f∥ * ∏ i, ∥p.apply_composition c v i∥ : continuous_multilinear_map.le_op_norm _ _ +... ≤ ∥f∥ * ∏ i, ∥p (c.blocks_fun i)∥ * + ∏ j : fin (c.blocks_fun i), ∥(v ∘ (c.embedding i)) j∥ : + begin + apply mul_le_mul_of_nonneg_left _ (norm_nonneg _), + refine finset.prod_le_prod (λ i hi, norm_nonneg _) (λ i hi, _), + apply continuous_multilinear_map.le_op_norm, + end +... = ∥f∥ * (∏ i, ∥p (c.blocks_fun i)∥) * + ∏ i (j : fin (c.blocks_fun i)), ∥(v ∘ (c.embedding i)) j∥ : + by rw [finset.prod_mul_distrib, mul_assoc] +... = ∥f∥ * (∏ i, ∥p (c.blocks_fun i)∥) * (∏ i : fin n, ∥v i∥) : + by { rw [← c.blocks_fin_equiv.prod_comp, ← finset.univ_sigma_univ, finset.prod_sigma], + congr } + +/-- The norm of `q.comp_along_composition p c` is controlled by the product of +the norms of the relevant bits of `q` and `p`. -/ +lemma comp_along_composition_norm {n : ℕ} + (q : formal_multilinear_series 𝕜 F G) (p : formal_multilinear_series 𝕜 E F) + (c : composition n) : + ∥q.comp_along_composition p c∥ ≤ ∥q c.length∥ * ∏ i, ∥p (c.blocks_fun i)∥ := +continuous_multilinear_map.op_norm_le_bound _ + (mul_nonneg (norm_nonneg _) (finset.prod_nonneg (λ i hi, norm_nonneg _))) + (comp_along_composition_bound _ _ _) + +lemma comp_along_composition_nnnorm {n : ℕ} + (q : formal_multilinear_series 𝕜 F G) (p : formal_multilinear_series 𝕜 E F) + (c : composition n) : + ∥q.comp_along_composition p c∥₊ ≤ ∥q c.length∥₊ * ∏ i, ∥p (c.blocks_fun i)∥₊ := +by { rw ← nnreal.coe_le_coe, push_cast, exact q.comp_along_composition_norm p c } + /-! ### The identity formal power series diff --git a/src/analysis/normed_space/indicator_function.lean b/src/analysis/normed_space/indicator_function.lean index 746e207f9b985..71019a37d71c5 100644 --- a/src/analysis/normed_space/indicator_function.lean +++ b/src/analysis/normed_space/indicator_function.lean @@ -15,7 +15,7 @@ This file contains a few simple lemmas about `set.indicator` and `norm`. indicator, norm -/ -variables {α E : Type*} [normed_group E] {s t : set α} (f : α → E) (a : α) +variables {α E : Type*} [semi_normed_group E] {s t : set α} (f : α → E) (a : α) open set From 90d6f2710f1463b34279fb1b79d497bcec4156d0 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 4 May 2022 17:58:32 +0000 Subject: [PATCH 114/153] ci(workflows/dependent-issues): run once every 15 mins, instead of on every merged PR (#13940) --- .github/workflows/dependent-issues.yml | 13 +------------ 1 file changed, 1 insertion(+), 12 deletions(-) diff --git a/.github/workflows/dependent-issues.yml b/.github/workflows/dependent-issues.yml index 48e6eace562e9..0efb065dc44c4 100644 --- a/.github/workflows/dependent-issues.yml +++ b/.github/workflows/dependent-issues.yml @@ -1,19 +1,8 @@ name: Dependent Issues on: - issues: - types: - - opened - - edited - - reopened - pull_request_target: - types: - - opened - - edited - - reopened - - synchronize schedule: - - cron: '15 * * * *' # schedule hourly check for external dependencies + - cron: '*/15 * * * *' # run every 15 minutes jobs: cancel: From 91c0ef86c522355b76acbfa92def7e9512144e72 Mon Sep 17 00:00:00 2001 From: kkytola Date: Wed, 4 May 2022 20:04:15 +0000 Subject: [PATCH 115/153] feat(analysis/normed_space/weak_dual): add the rest of Banach-Alaoglu theorem (#9862) The second of two parts to add the Banach-Alaoglu theorem about the compactness of the closed unit ball (and more generally polar sets of neighborhoods of the origin) of the dual of a normed space in the weak-star topology. This second half is about the embedding of the weak dual of a normed space into a (big) product of the ground fields, and the required compactness statements from Tychonoff's theorem. In particular it contains the actual Banach-Alaoglu theorem. Co-Authored-By: Yury Kudryashov Co-authored-by: kkytola <39528102+kkytola@users.noreply.github.com> --- src/analysis/normed_space/dual.lean | 4 +- src/analysis/normed_space/operator_norm.lean | 76 ++++++++- src/analysis/normed_space/weak_dual.lean | 161 +++++++++++++------ src/topology/algebra/module/weak_dual.lean | 6 +- src/topology/maps.lean | 4 + src/topology/order.lean | 7 +- 6 files changed, 205 insertions(+), 53 deletions(-) diff --git a/src/analysis/normed_space/dual.lean b/src/analysis/normed_space/dual.lean index b9987efac9c29..7127fe14ad1ce 100644 --- a/src/analysis/normed_space/dual.lean +++ b/src/analysis/normed_space/dual.lean @@ -142,11 +142,11 @@ open metric set normed_space `polar 𝕜 s` is the subset of `dual 𝕜 E` consisting of those functionals which evaluate to something of norm at most one at all points `z ∈ s`. -/ def polar (𝕜 : Type*) [nondiscrete_normed_field 𝕜] - {E : Type*} [normed_group E] [normed_space 𝕜 E] : set E → set (dual 𝕜 E) := + {E : Type*} [semi_normed_group E] [normed_space 𝕜 E] : set E → set (dual 𝕜 E) := (dual_pairing 𝕜 E).flip.polar variables (𝕜 : Type*) [nondiscrete_normed_field 𝕜] -variables {E : Type*} [normed_group E] [normed_space 𝕜 E] +variables {E : Type*} [semi_normed_group E] [normed_space 𝕜 E] lemma mem_polar_iff {x' : dual 𝕜 E} (s : set E) : x' ∈ polar 𝕜 s ↔ ∀ z ∈ s, ∥x' z∥ ≤ 1 := iff.rfl diff --git a/src/analysis/normed_space/operator_norm.lean b/src/analysis/normed_space/operator_norm.lean index e0090f6bfa77f..487928749e3cc 100644 --- a/src/analysis/normed_space/operator_norm.lean +++ b/src/analysis/normed_space/operator_norm.lean @@ -1357,7 +1357,7 @@ that it belongs to the closure of the image of a bounded set `s : set (E →SL[ to function. Coercion to function of the result is definitionally equal to `f`. -/ @[simps apply { fully_applied := ff }] def of_mem_closure_image_coe_bounded (f : E' → F) {s : set (E' →SL[σ₁₂] F)} (hs : bounded s) - (hf : f ∈ closure ((λ g x, g x : (E' →SL[σ₁₂] F) → E' → F) '' s)) : + (hf : f ∈ closure ((coe_fn : (E' →SL[σ₁₂] F) → E' → F) '' s)) : E' →SL[σ₁₂] F := begin -- `f` is a linear map due to `linear_map_of_mem_closure_range_coe` @@ -1428,6 +1428,80 @@ begin exact ⟨Glin, tendsto_of_tendsto_pointwise_of_cauchy_seq (tendsto_pi_nhds.2 hG) hf⟩ end +/-- Let `s` be a bounded set in the space of continuous (semi)linear maps `E →SL[σ] F` taking values +in a proper space. Then `s` interpreted as a set in the space of maps `E → F` with topology of +pointwise convergence is precompact: its closure is a compact set. -/ +lemma is_compact_closure_image_coe_of_bounded [proper_space F] {s : set (E' →SL[σ₁₂] F)} + (hb : bounded s) : + is_compact (closure ((coe_fn : (E' →SL[σ₁₂] F) → E' → F) '' s)) := +have ∀ x, is_compact (closure (apply' F σ₁₂ x '' s)), + from λ x, ((apply' F σ₁₂ x).lipschitz.bounded_image hb).is_compact_closure, +compact_closure_of_subset_compact (is_compact_pi_infinite this) + (image_subset_iff.2 $ λ g hg x, subset_closure $ mem_image_of_mem _ hg) + +/-- Let `s` be a bounded set in the space of continuous (semi)linear maps `E →SL[σ] F` taking values +in a proper space. If `s` interpreted as a set in the space of maps `E → F` with topology of +pointwise convergence is closed, then it is compact. + +TODO: reformulate this in terms of a type synonym with the right topology. -/ +lemma is_compact_image_coe_of_bounded_of_closed_image [proper_space F] {s : set (E' →SL[σ₁₂] F)} + (hb : bounded s) (hc : is_closed ((coe_fn : (E' →SL[σ₁₂] F) → E' → F) '' s)) : + is_compact ((coe_fn : (E' →SL[σ₁₂] F) → E' → F) '' s) := +hc.closure_eq ▸ is_compact_closure_image_coe_of_bounded hb + +/-- If a set `s` of semilinear functions is bounded and is closed in the weak-* topology, then its +image under coercion to functions `E → F` is a closed set. We don't have a name for `E →SL[σ] F` +with weak-* topology in `mathlib`, so we use an equivalent condition (see `is_closed_induced_iff'`). + +TODO: reformulate this in terms of a type synonym with the right topology. -/ +lemma is_closed_image_coe_of_bounded_of_weak_closed {s : set (E' →SL[σ₁₂] F)} (hb : bounded s) + (hc : ∀ f, (⇑f : E' → F) ∈ closure ((coe_fn : (E' →SL[σ₁₂] F) → E' → F) '' s) → f ∈ s) : + is_closed ((coe_fn : (E' →SL[σ₁₂] F) → E' → F) '' s) := +is_closed_of_closure_subset $ λ f hf, + ⟨of_mem_closure_image_coe_bounded f hb hf, hc (of_mem_closure_image_coe_bounded f hb hf) hf, rfl⟩ + +/-- If a set `s` of semilinear functions is bounded and is closed in the weak-* topology, then its +image under coercion to functions `E → F` is a compact set. We don't have a name for `E →SL[σ] F` +with weak-* topology in `mathlib`, so we use an equivalent condition (see `is_closed_induced_iff'`). +-/ +lemma is_compact_image_coe_of_bounded_of_weak_closed [proper_space F] {s : set (E' →SL[σ₁₂] F)} + (hb : bounded s) + (hc : ∀ f, (⇑f : E' → F) ∈ closure ((coe_fn : (E' →SL[σ₁₂] F) → E' → F) '' s) → f ∈ s) : + is_compact ((coe_fn : (E' →SL[σ₁₂] F) → E' → F) '' s) := +is_compact_image_coe_of_bounded_of_closed_image hb $ + is_closed_image_coe_of_bounded_of_weak_closed hb hc + +/-- A closed ball is closed in the weak-* topology. We don't have a name for `E →SL[σ] F` with +weak-* topology in `mathlib`, so we use an equivalent condition (see `is_closed_induced_iff'`). -/ +lemma is_weak_closed_closed_ball (f₀ : E' →SL[σ₁₂] F) (r : ℝ) ⦃f : E' →SL[σ₁₂] F⦄ + (hf : ⇑f ∈ closure ((coe_fn : (E' →SL[σ₁₂] F) → E' → F) '' (closed_ball f₀ r))) : + f ∈ closed_ball f₀ r := +begin + have hr : 0 ≤ r, + from nonempty_closed_ball.1 (nonempty_image_iff.1 (closure_nonempty_iff.1 ⟨_, hf⟩)), + refine mem_closed_ball_iff_norm.2 (op_norm_le_bound _ hr $ λ x, _), + have : is_closed {g : E' → F | ∥g x - f₀ x∥ ≤ r * ∥x∥}, + from is_closed_Iic.preimage ((@continuous_apply E' (λ _, F) _ x).sub continuous_const).norm, + refine this.closure_subset_iff.2 (image_subset_iff.2 $ λ g hg, _) hf, + exact (g - f₀).le_of_op_norm_le (mem_closed_ball_iff_norm.1 hg) _ +end + +/-- The set of functions `f : E → F` that represent continuous linear maps `f : E →SL[σ₁₂] F` +at distance `≤ r` from `f₀ : E →SL[σ₁₂] F` is closed in the topology of pointwise convergence. +This is one of the key steps in the proof of the **Banach-Alaoglu** theorem. -/ +lemma is_closed_image_coe_closed_ball (f₀ : E →SL[σ₁₂] F) (r : ℝ) : + is_closed ((coe_fn : (E →SL[σ₁₂] F) → E → F) '' closed_ball f₀ r) := +is_closed_image_coe_of_bounded_of_weak_closed bounded_closed_ball (is_weak_closed_closed_ball f₀ r) + +/-- **Banach-Alaoglu** theorem. The set of functions `f : E → F` that represent continuous linear +maps `f : E →SL[σ₁₂] F` at distance `≤ r` from `f₀ : E →SL[σ₁₂] F` is compact in the topology of +pointwise convergence. Other versions of this theorem can be found in +`analysis.normed_space.weak_dual`. -/ +lemma is_compact_image_coe_closed_ball [proper_space F] (f₀ : E →SL[σ₁₂] F) (r : ℝ) : + is_compact ((coe_fn : (E →SL[σ₁₂] F) → E → F) '' closed_ball f₀ r) := +is_compact_image_coe_of_bounded_of_weak_closed bounded_closed_ball $ + is_weak_closed_closed_ball f₀ r + end completeness section uniformly_extend diff --git a/src/analysis/normed_space/weak_dual.lean b/src/analysis/normed_space/weak_dual.lean index 4bac47c508a80..f405d40bf1078 100644 --- a/src/analysis/normed_space/weak_dual.lean +++ b/src/analysis/normed_space/weak_dual.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2021 Kalle Kytölä. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Kalle Kytölä +Authors: Kalle Kytölä, Yury Kudryashov -/ import topology.algebra.module.weak_dual import analysis.normed_space.dual @@ -19,7 +19,9 @@ It is shown that the canonical mapping `normed_space.dual 𝕜 E → weak_dual as a consequence the weak-* topology is coarser than the topology obtained from the operator norm (dual norm). -The file is a stub, some TODOs below. +In this file, we also establish the Banach-Alaoglu theorem about the compactness of closed balls +in the dual of `E` (as well as sets of somewhat more general form) with respect to the weak-* +topology. ## Main definitions @@ -37,15 +39,21 @@ The first main result concerns the comparison of the operator norm topology on ` weak-* topology on (its type synonym) `weak_dual 𝕜 E`: * `dual_norm_topology_le_weak_dual_topology`: The weak-* topology on the dual of a normed space is coarser (not necessarily strictly) than the operator norm topology. +* `weak_dual.is_compact_polar` (a version of the Banach-Alaoglu theorem): The polar set of a + neighborhood of the origin in a normed space `E` over `𝕜` is compact in `weak_dual _ E`, if the + nondiscrete normed field `𝕜` is proper as a topological space. +* `weak_dual.is_compact_closed_ball` (the most common special case of the Banach-Alaoglu theorem): + Closed balls in the dual of a normed space `E` over `ℝ` or `ℂ` are compact in the weak-star + topology. TODOs: * Add that in finite dimensions, the weak-* topology and the dual norm topology coincide. * Add that in infinite dimensions, the weak-* topology is strictly coarser than the dual norm topology. -* Add Banach-Alaoglu theorem (general version maybe in `topology.algebra.module.weak_dual`). -* Add metrizability of the dual unit ball (more generally bounded subsets) of `weak_dual 𝕜 E` - under the assumption of separability of `E`. Sequential Banach-Alaoglu theorem would then follow - from the general one. +* Add metrizability of the dual unit ball (more generally weak-star compact subsets) of + `weak_dual 𝕜 E` under the assumption of separability of `E`. +* Add the sequential Banach-Alaoglu theorem: the dual unit ball of a separable normed space `E` + is sequentially compact in the weak-star topology. This would follow from the metrizability above. ## Notations @@ -58,9 +66,18 @@ Weak-* topology is defined generally in the file `topology.algebra.module.weak_d When `E` is a normed space, the duals `dual 𝕜 E` and `weak_dual 𝕜 E` are type synonyms with different topology instances. +For the proof of Banach-Alaoglu theorem, the weak dual of `E` is embedded in the space of +functions `E → 𝕜` with the topology of pointwise convergence. + +The polar set `polar 𝕜 s` of a subset `s` of `E` is originally defined as a subset of the dual +`dual 𝕜 E`. We care about properties of these w.r.t. weak-* topology, and for this purpose give +the definition `weak_dual.polar 𝕜 s` for the "same" subset viewed as a subset of `weak_dual 𝕜 E` +(a type synonym of the dual but with a different topology instance). + ## References * https://en.wikipedia.org/wiki/Weak_topology#Weak-*_topology +* https://en.wikipedia.org/wiki/Banach%E2%80%93Alaoglu_theorem ## Tags @@ -69,53 +86,37 @@ weak-star, weak dual -/ noncomputable theory -open filter -open_locale topological_space +open filter function metric set +open_locale topological_space filter /-! ### Weak star topology on duals of normed spaces + In this section, we prove properties about the weak-* topology on duals of normed spaces. We prove in particular that the canonical mapping `dual 𝕜 E → weak_dual 𝕜 E` is continuous, i.e., that the weak-* topology is coarser (not necessarily strictly) than the topology given by the dual-norm (i.e. the operator-norm). -/ -open normed_space - variables {𝕜 : Type*} [nondiscrete_normed_field 𝕜] variables {E : Type*} [semi_normed_group E] [normed_space 𝕜 E] -/-- For normed spaces `E`, there is a canonical map `dual 𝕜 E → weak_dual 𝕜 E` (the "identity" -mapping). It is a linear equivalence. -/ -def normed_space.dual.to_weak_dual : dual 𝕜 E ≃ₗ[𝕜] weak_dual 𝕜 E := -linear_equiv.refl 𝕜 (E →L[𝕜] 𝕜) +namespace normed_space -/-- For normed spaces `E`, there is a canonical map `weak_dual 𝕜 E → dual 𝕜 E` (the "identity" -mapping). It is a linear equivalence. Here it is implemented as the inverse of the linear -equivalence `normed_space.dual.to_weak_dual` in the other direction. -/ -def weak_dual.to_normed_dual : weak_dual 𝕜 E ≃ₗ[𝕜] dual 𝕜 E := -normed_space.dual.to_weak_dual.symm +namespace dual -@[simp] lemma weak_dual.coe_to_fun_eq_normed_coe_to_fun (x' : dual 𝕜 E) : - (x'.to_weak_dual : E → 𝕜) = x' := rfl +/-- For normed spaces `E`, there is a canonical map `dual 𝕜 E → weak_dual 𝕜 E` (the "identity" +mapping). It is a linear equivalence. -/ +def to_weak_dual : dual 𝕜 E ≃ₗ[𝕜] weak_dual 𝕜 E := linear_equiv.refl 𝕜 (E →L[𝕜] 𝕜) -namespace normed_space.dual +@[simp] lemma coe_to_weak_dual (x' : dual 𝕜 E) : ⇑(x'.to_weak_dual) = x' := rfl @[simp] lemma to_weak_dual_eq_iff (x' y' : dual 𝕜 E) : x'.to_weak_dual = y'.to_weak_dual ↔ x' = y' := to_weak_dual.injective.eq_iff -@[simp] lemma _root_.weak_dual.to_normed_dual_eq_iff (x' y' : weak_dual 𝕜 E) : - x'.to_normed_dual = y'.to_normed_dual ↔ x' = y' := -weak_dual.to_normed_dual.injective.eq_iff - -theorem to_weak_dual_continuous : - continuous (λ (x' : dual 𝕜 E), x'.to_weak_dual) := -begin - apply continuous_of_continuous_eval, - intros z, - exact (inclusion_in_double_dual 𝕜 E z).continuous, -end +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 /-- 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 @@ -127,25 +128,91 @@ def continuous_linear_map_to_weak_dual : dual 𝕜 E →L[𝕜] weak_dual 𝕜 E theorem dual_norm_topology_le_weak_dual_topology : (by apply_instance : topological_space (dual 𝕜 E)) ≤ (by apply_instance : topological_space (weak_dual 𝕜 E)) := -begin - refine continuous.le_induced _, - apply continuous_pi_iff.mpr, - intros z, - exact (inclusion_in_double_dual 𝕜 E z).continuous, -end +by { convert to_weak_dual_continuous.le_induced, exact induced_id.symm } -end normed_space.dual +end dual +end normed_space namespace weak_dual -lemma to_normed_dual.preimage_closed_unit_ball : - (to_normed_dual ⁻¹' metric.closed_ball (0 : dual 𝕜 E) 1) = - {x' : weak_dual 𝕜 E | ∥ x'.to_normed_dual ∥ ≤ 1} := +open normed_space + +/-- For normed spaces `E`, there is a canonical map `weak_dual 𝕜 E → dual 𝕜 E` (the "identity" +mapping). It is a linear equivalence. Here it is implemented as the inverse of the linear +equivalence `normed_space.dual.to_weak_dual` in the other direction. -/ +def to_normed_dual : weak_dual 𝕜 E ≃ₗ[𝕜] dual 𝕜 E := normed_space.dual.to_weak_dual.symm + +@[simp] lemma coe_to_normed_dual (x' : weak_dual 𝕜 E) : ⇑(x'.to_normed_dual) = x' := rfl + +@[simp] lemma to_normed_dual_eq_iff (x' y' : weak_dual 𝕜 E) : + x'.to_normed_dual = y'.to_normed_dual ↔ x' = y' := +weak_dual.to_normed_dual.injective.eq_iff + +lemma is_closed_closed_ball (x' : dual 𝕜 E) (r : ℝ) : + is_closed (to_normed_dual ⁻¹' closed_ball x' r) := +is_closed_induced_iff'.2 (continuous_linear_map.is_weak_closed_closed_ball x' r) + +/-! +### Polar sets in the weak dual space +-/ + +variables (𝕜) + +/-- The polar set `polar 𝕜 s` of `s : set E` seen as a subset of the dual of `E` with the +weak-star topology is `weak_dual.polar 𝕜 s`. -/ +def polar (s : set E) : set (weak_dual 𝕜 E) := to_normed_dual ⁻¹' polar 𝕜 s + +lemma polar_def (s : set E) : polar 𝕜 s = {f : weak_dual 𝕜 E | ∀ x ∈ s, ∥f x∥ ≤ 1} := rfl + +/-- The polar `polar 𝕜 s` of a set `s : E` is a closed subset when the weak star topology +is used. -/ +lemma is_closed_polar (s : set E) : is_closed (polar 𝕜 s) := begin - have eq : metric.closed_ball (0 : dual 𝕜 E) 1 = {x' : dual 𝕜 E | ∥ x' ∥ ≤ 1}, - { ext x', simp only [dist_zero_right, metric.mem_closed_ball, set.mem_set_of_eq], }, - rw eq, - exact set.preimage_set_of_eq, + simp only [polar_def, set_of_forall], + exact is_closed_bInter (λ x hx, is_closed_Iic.preimage (eval_continuous _ _).norm) end +variable {𝕜} + +/-- While the coercion `coe_fn : weak_dual 𝕜 E → (E → 𝕜)` is not a closed map, it sends *bounded* +closed sets to closed sets. -/ +lemma is_closed_image_coe_of_bounded_of_closed {s : set (weak_dual 𝕜 E)} + (hb : bounded (dual.to_weak_dual ⁻¹' s)) (hc : is_closed s) : + is_closed ((coe_fn : weak_dual 𝕜 E → E → 𝕜) '' s) := +continuous_linear_map.is_closed_image_coe_of_bounded_of_weak_closed hb (is_closed_induced_iff'.1 hc) + +lemma is_compact_of_bounded_of_closed [proper_space 𝕜] {s : set (weak_dual 𝕜 E)} + (hb : bounded (dual.to_weak_dual ⁻¹' s)) (hc : is_closed s) : + is_compact s := +(embedding.is_compact_iff_is_compact_image fun_like.coe_injective.embedding_induced).mpr $ + continuous_linear_map.is_compact_image_coe_of_bounded_of_closed_image hb $ + is_closed_image_coe_of_bounded_of_closed hb hc + +variable (𝕜) + +/-- The image under `coe_fn : weak_dual 𝕜 E → (E → 𝕜)` of a polar `weak_dual.polar 𝕜 s` of a +neighborhood `s` of the origin is a closed set. -/ +lemma is_closed_image_polar_of_mem_nhds {s : set E} (s_nhd : s ∈ 𝓝 (0 : E)) : + is_closed ((coe_fn : weak_dual 𝕜 E → E → 𝕜) '' polar 𝕜 s) := +is_closed_image_coe_of_bounded_of_closed (bounded_polar_of_mem_nhds_zero 𝕜 s_nhd) + (is_closed_polar _ _) + +/-- The image under `coe_fn : normed_space.dual 𝕜 E → (E → 𝕜)` of a polar `polar 𝕜 s` of a +neighborhood `s` of the origin is a closed set. -/ +lemma _root_.normed_space.dual.is_closed_image_polar_of_mem_nhds {s : set E} + (s_nhd : s ∈ 𝓝 (0 : E)) : is_closed ((coe_fn : dual 𝕜 E → E → 𝕜) '' normed_space.polar 𝕜 s) := +is_closed_image_polar_of_mem_nhds 𝕜 s_nhd + +/-- The **Banach-Alaoglu theorem**: the polar set of a neighborhood `s` of the origin in a +normed space `E` is a compact subset of `weak_dual 𝕜 E`. -/ +theorem is_compact_polar [proper_space 𝕜] {s : set E} (s_nhd : s ∈ 𝓝 (0 : E)) : + is_compact (polar 𝕜 s) := +is_compact_of_bounded_of_closed (bounded_polar_of_mem_nhds_zero 𝕜 s_nhd) (is_closed_polar _ _) + +/-- The **Banach-Alaoglu theorem**: closed balls of the dual of a normed space `E` are compact in +the weak-star topology. -/ +theorem is_compact_closed_ball [proper_space 𝕜] (x' : dual 𝕜 E) (r : ℝ) : + is_compact (to_normed_dual ⁻¹' (closed_ball x' r)) := +is_compact_of_bounded_of_closed bounded_closed_ball (is_closed_closed_ball x' r) + end weak_dual diff --git a/src/topology/algebra/module/weak_dual.lean b/src/topology/algebra/module/weak_dual.lean index 8adef9c5d0840..10c9126b6706f 100644 --- a/src/topology/algebra/module/weak_dual.lean +++ b/src/topology/algebra/module/weak_dual.lean @@ -180,9 +180,13 @@ weak_bilin (top_dual_pairing 𝕜 E) instance : inhabited (weak_dual 𝕜 E) := continuous_linear_map.inhabited -instance add_monoid_hom_class_weak_dual : add_monoid_hom_class (weak_dual 𝕜 E) E 𝕜 := +instance weak_dual.add_monoid_hom_class : add_monoid_hom_class (weak_dual 𝕜 E) E 𝕜 := continuous_linear_map.add_monoid_hom_class +/-- Helper instance for when there's too many metavariables to apply `fun_like.has_coe_to_fun` +directly. -/ +instance : has_coe_to_fun (weak_dual 𝕜 E) (λ _, E → 𝕜) := fun_like.has_coe_to_fun + /-- If a monoid `M` distributively continuously acts on `𝕜` and this action commutes with multiplication on `𝕜`, then it acts on `weak_dual 𝕜 E`. -/ instance (M) [monoid M] [distrib_mul_action M 𝕜] [smul_comm_class 𝕜 M 𝕜] diff --git a/src/topology/maps.lean b/src/topology/maps.lean index 7cd38018c8dd1..3dcbb381db248 100644 --- a/src/topology/maps.lean +++ b/src/topology/maps.lean @@ -121,6 +121,10 @@ lemma inducing.is_closed_iff {f : α → β} (hf : inducing f) {s : set α} : is_closed s ↔ ∃ t, is_closed t ∧ f ⁻¹' t = s := by rw [hf.induced, is_closed_induced_iff] +lemma inducing.is_closed_iff' {f : α → β} (hf : inducing f) {s : set α} : + is_closed s ↔ ∀ x, f x ∈ closure (f '' s) → x ∈ s := +by rw [hf.induced, is_closed_induced_iff'] + lemma inducing.is_open_iff {f : α → β} (hf : inducing f) {s : set α} : is_open s ↔ ∃ t, is_open t ∧ f ⁻¹' t = s := by rw [hf.induced, is_open_induced_iff] diff --git a/src/topology/order.lean b/src/topology/order.lean index ede00f242234c..8ceb8536a921a 100644 --- a/src/topology/order.lean +++ b/src/topology/order.lean @@ -330,8 +330,7 @@ lemma is_closed_induced_iff [t : topological_space β] {s : set α} {f : α → @is_closed α (t.induced f) s ↔ (∃t, is_closed t ∧ f ⁻¹' t = s) := begin simp only [← is_open_compl_iff, is_open_induced_iff], - exact ⟨λ ⟨t, ht, heq⟩, ⟨tᶜ, by rwa compl_compl, by simp [preimage_compl, heq, compl_compl]⟩, - λ ⟨t, ht, heq⟩, ⟨tᶜ, ht, by simp only [preimage_compl, heq.symm]⟩⟩ + exact compl_surjective.exists.trans (by simp only [preimage_compl, compl_inj_iff]) end /-- Given `f : α → β` and a topology on `α`, the coinduced topology on `β` is defined @@ -780,6 +779,10 @@ lemma closure_induced [t : topological_space β] {f : α → β} {a : α} {s : s a ∈ @closure α (topological_space.induced f t) s ↔ f a ∈ closure (f '' s) := by simp only [mem_closure_iff_frequently, nhds_induced, frequently_comap, mem_image, and_comm] +lemma is_closed_induced_iff' [t : topological_space β] {f : α → β} {s : set α} : + @is_closed α (t.induced f) s ↔ ∀ a, f a ∈ closure (f '' s) → a ∈ s := +by simp only [← closure_subset_iff_is_closed, subset_def, closure_induced] + end induced section sierpinski From fd8474f55263f48c9f5293b8cf6ece4737145cf2 Mon Sep 17 00:00:00 2001 From: Christopher Hoskin Date: Wed, 4 May 2022 20:04:17 +0000 Subject: [PATCH 116/153] feat(algebra/ring/idempotents): Introduce idempotents (#13830) Co-authored-by: Christopher Hoskin --- src/algebra/ring/idempotents.lean | 127 ++++++++++++++++++++++++++++++ 1 file changed, 127 insertions(+) create mode 100644 src/algebra/ring/idempotents.lean diff --git a/src/algebra/ring/idempotents.lean b/src/algebra/ring/idempotents.lean new file mode 100644 index 0000000000000..26a0130cbe0e8 --- /dev/null +++ b/src/algebra/ring/idempotents.lean @@ -0,0 +1,127 @@ +/- +Copyright (c) 2022 Christopher Hoskin. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Christopher Hoskin +-/ +import algebra.ring.basic +import algebra.group_power.basic +import tactic.nth_rewrite.default + +/-! +# Idempotents + +This file defines idempotents for an arbitary multiplication and proves some basic results, +including: + +* `is_idempotent_elem.mul_of_commute`: In a semigroup, the product of two commuting idempotents is + an idempotent; +* `is_idempotent_elem.one_sub_iff`: In a (non-associative) ring, `p` is an idempotent if and only if + `1-p` is an idempotent. +* `is_idempotent_elem.pow_succ_eq`: In a monoid `p ^ (n+1) = p` for `p` an idempotent and `n` a + natural number. + +## Tags + +projection, idempotent +-/ + +variables {M N S M₀ M₁ R G G₀ : Type*} +variables [has_mul M] [monoid N] [semigroup S] [mul_zero_class M₀] [mul_one_class M₁] + [non_assoc_ring R] [group G] [group_with_zero G₀] + +/-- +An element `p` is said to be idempotent if `p * p = p` +-/ +def is_idempotent_elem (p : M) : Prop := p * p = p + +namespace is_idempotent_elem + +lemma of_is_idempotent [is_idempotent M (*)] (a : M) : is_idempotent_elem a := +is_idempotent.idempotent a + +lemma eq {p : M} (h : is_idempotent_elem p) : p * p = p := h + +lemma mul_of_commute {p q : S} (h : commute p q) (h₁ : is_idempotent_elem p) + (h₂ : is_idempotent_elem q) : is_idempotent_elem (p * q) := + by rw [is_idempotent_elem, mul_assoc, ← mul_assoc q, ←h.eq, mul_assoc p, h₂.eq, ← mul_assoc, + h₁.eq] + +lemma zero : is_idempotent_elem (0 : M₀) := mul_zero _ + +lemma one : is_idempotent_elem (1 : M₁) := mul_one _ + +lemma one_sub {p : R} (h : is_idempotent_elem p) : is_idempotent_elem (1 - p) := +begin + rw is_idempotent_elem at h, + rw [is_idempotent_elem, mul_sub_left_distrib, mul_one, sub_mul, one_mul, h, sub_self, sub_zero], +end + +@[simp] lemma one_sub_iff {p : R} : is_idempotent_elem (1 - p) ↔ is_idempotent_elem p := +⟨ λ h, sub_sub_cancel 1 p ▸ h.one_sub, is_idempotent_elem.one_sub ⟩ + +lemma pow {p : N} (n : ℕ) (h : is_idempotent_elem p) : is_idempotent_elem (p ^ n) := +begin + induction n with n ih, + { rw pow_zero, apply one, }, + { unfold is_idempotent_elem, + rw [pow_succ, ← mul_assoc, ← pow_mul_comm', mul_assoc (p^n), h.eq, pow_mul_comm', mul_assoc, + ih.eq], } +end + +lemma pow_succ_eq {p : N} (n : ℕ) (h : is_idempotent_elem p) : p ^ (n + 1) = p := +begin + induction n with n ih, + { rw [nat.zero_add, pow_one], }, + { rw [pow_succ, ih, h.eq], } +end + +@[simp] lemma iff_eq_one {p : G} : is_idempotent_elem p ↔ p = 1 := +begin + split, + { intro h, + rw ← mul_left_inv p, + nth_rewrite_rhs 1 ← h.eq, + rw [← mul_assoc, mul_left_inv, one_mul] }, + { intro h, rw h, apply one, } +end + +@[simp] lemma iff_eq_zero_or_one {p : G₀} : is_idempotent_elem p ↔ p = 0 ∨ p = 1 := +begin + split, + { intro h, + by_cases hp : p = 0, + { finish, }, + { rw ← mul_inv_cancel hp, + nth_rewrite_rhs 1 ← h.eq, + rw [mul_assoc, mul_inv_cancel hp, mul_one, eq_self_iff_true, or_true], + finish, } }, + { rintro (h₁ | h₂), + { rw h₁, apply zero, }, + { rw h₂, apply one, } } +end + +/-! ### Instances on `subtype is_idempotent_elem` -/ +section instances + +instance : has_zero { p : M₀ // is_idempotent_elem p } := { zero := ⟨ 0, zero ⟩ } + +@[simp] lemma coe_zero : ↑(0 : {p : M₀ // is_idempotent_elem p}) = (0 : M₀) := rfl + +instance : has_one { p : M₁ // is_idempotent_elem p } := { one := ⟨ 1, one ⟩ } + +@[simp] lemma coe_one : ↑(1 : { p : M₁ // is_idempotent_elem p }) = (1 : M₁) := rfl + +instance : has_compl { p : R // is_idempotent_elem p } := ⟨λ p, ⟨1 - p, p.prop.one_sub⟩⟩ + +@[simp] lemma coe_compl (p : { p : R // is_idempotent_elem p }) : ↑(pᶜ) = (1 : R) - ↑p := rfl + +@[simp] lemma compl_compl (p : {p : R // is_idempotent_elem p}) : pᶜᶜ = p := +subtype.ext $ sub_sub_cancel _ _ + +@[simp] lemma zero_compl : (0 : {p : R // is_idempotent_elem p})ᶜ = 1 := subtype.ext $ sub_zero _ + +@[simp] lemma one_compl : (1 : {p : R // is_idempotent_elem p})ᶜ = 0 := subtype.ext $ sub_self _ + +end instances + +end is_idempotent_elem From edf6cefaa1ab4eb886c632102f62a87691647724 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 4 May 2022 20:04:18 +0000 Subject: [PATCH 117/153] feat(set_theory/game/nim): `nim 0` is a relabelling of `0` and `nim 1` is a relabelling of `star` (#13846) --- src/set_theory/game/nim.lean | 38 ++++++++++++++++++++++++-- src/set_theory/ordinal/arithmetic.lean | 21 +++++++++++--- 2 files changed, 52 insertions(+), 7 deletions(-) diff --git a/src/set_theory/game/nim.lean b/src/set_theory/game/nim.lean index cf8bb340f12f5..197ef5036e57d 100644 --- a/src/set_theory/game/nim.lean +++ b/src/set_theory/game/nim.lean @@ -57,6 +57,35 @@ lemma nim_def (O : ordinal) : nim O = pgame.mk O.out.α O.out.α (λ O₂, nim (ordinal.typein (<) O₂)) := by { rw nim, refl } +instance : is_empty (left_moves (nim 0)) := +by { rw nim_def, exact α.is_empty } + +instance : is_empty (right_moves (nim 0)) := +by { rw nim_def, exact α.is_empty } + +noncomputable instance : unique (left_moves (nim 1)) := +by { rw nim_def, exact α.unique } + +noncomputable instance : unique (right_moves (nim 1)) := +by { rw nim_def, exact α.unique } + +/-- `0` has exactly the same moves as `nim 0`. -/ +def nim_zero_relabelling : relabelling 0 (nim 0) := +(relabelling.is_empty _).symm + +@[simp] theorem nim_zero_equiv : 0 ≈ nim 0 := nim_zero_relabelling.equiv + +/-- `nim 1` has exactly the same moves as `star`. -/ +noncomputable def nim_one_relabelling : relabelling star (nim 1) := +begin + rw nim_def, + refine ⟨_, _, λ i, _, λ j, _⟩, + any_goals { dsimp, apply equiv_of_unique_of_unique }, + all_goals { simp, exact nim_zero_relabelling } +end + +@[simp] theorem nim_one_equiv : star ≈ nim 1 := nim_one_relabelling.equiv + @[simp] lemma nim_birthday (O : ordinal) : (nim O).birthday = O := begin induction O using ordinal.induction with O IH, @@ -100,7 +129,7 @@ lemma non_zero_first_wins {O : ordinal} (hO : O ≠ 0) : (nim O).first_wins := begin rw [impartial.first_wins_symm, nim_def, lt_def_le], rw ←ordinal.pos_iff_ne_zero at hO, - exact or.inr ⟨(ordinal.principal_seg_out hO).top, by simpa using zero_first_loses.1⟩ + exact or.inr ⟨(ordinal.principal_seg_out hO).top, by simp⟩ end @[simp] lemma sum_first_loses_iff_eq (O₁ O₂ : ordinal) : (nim O₁ + nim O₂).first_loses ↔ O₁ = O₂ := @@ -194,12 +223,15 @@ by simp grundy_value G = grundy_value H ↔ G ≈ H := (grundy_value_eq_iff_equiv_nim _ _).trans (equiv_congr_left.1 (equiv_nim_grundy_value H) _).symm -@[simp] lemma grundy_value_zero : grundy_value 0 = 0 := -by rw [(grundy_value_eq_iff_equiv 0 (nim 0)).2 (equiv_symm nim.zero_first_loses), nim.grundy_value] +lemma grundy_value_zero : grundy_value 0 = 0 := +by simp @[simp] lemma grundy_value_iff_equiv_zero (G : pgame) [G.impartial] : grundy_value G = 0 ↔ G ≈ 0 := by rw [←grundy_value_eq_iff_equiv, grundy_value_zero] +lemma grundy_value_star : grundy_value star = 1 := +by simp + @[simp] lemma grundy_value_nim_add_nim (n m : ℕ) : grundy_value (nim.{u} n + nim.{u} m) = nat.lxor n m := begin diff --git a/src/set_theory/ordinal/arithmetic.lean b/src/set_theory/ordinal/arithmetic.lean index 46c56dc9097cd..dd84a0e9ccb43 100644 --- a/src/set_theory/ordinal/arithmetic.lean +++ b/src/set_theory/ordinal/arithmetic.lean @@ -176,12 +176,25 @@ type_ne_zero_iff_nonempty.2 ⟨punit.star⟩ instance : nontrivial ordinal.{u} := ⟨⟨1, 0, ordinal.one_ne_zero⟩⟩ -instance : nonempty (1 : ordinal).out.α := -out_nonempty_iff_ne_zero.2 ordinal.one_ne_zero - -theorem zero_lt_one : (0 : ordinal) < 1 := +@[simp] theorem zero_lt_one : (0 : ordinal) < 1 := lt_iff_le_and_ne.2 ⟨ordinal.zero_le _, ne.symm $ ordinal.one_ne_zero⟩ +instance : unique (1 : ordinal).out.α := +{ default := enum (<) 0 (by simp), + uniq := λ a, begin + rw ←enum_typein (<) a, + unfold default, + congr, + rw ←lt_one_iff_zero, + apply typein_lt_self + end } + +theorem one_out_eq (x : (1 : ordinal).out.α) : x = enum (<) 0 (by simp) := +unique.eq_default x + +@[simp] theorem typein_one_out (x : (1 : ordinal).out.α) : typein (<) x = 0 := +by rw [one_out_eq x, typein_enum] + theorem le_one_iff {a : ordinal} : a ≤ 1 ↔ a = 0 ∨ a = 1 := begin refine ⟨λ ha, _, _⟩, From a80e5683d1607bf62d580d4b26b80670e4239c1d Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 4 May 2022 20:04:19 +0000 Subject: [PATCH 118/153] feat(logic/equiv/set): equivalences between all preimages gives an equivalence of domains (#13853) Co-authored-by: Scott Morrison --- src/data/fintype/card.lean | 2 +- src/group_theory/coset.lean | 2 +- src/group_theory/group_action/basic.lean | 2 +- src/group_theory/p_group.lean | 2 +- src/logic/equiv/basic.lean | 26 +++++++++++++++++------- src/logic/equiv/set.lean | 17 ++++++++++++++++ src/logic/small.lean | 2 +- src/set_theory/cardinal/basic.lean | 4 ++-- 8 files changed, 43 insertions(+), 14 deletions(-) diff --git a/src/data/fintype/card.lean b/src/data/fintype/card.lean index 9d5e5a4f1493d..3faa0026a84c4 100644 --- a/src/data/fintype/card.lean +++ b/src/data/fintype/card.lean @@ -213,7 +213,7 @@ lemma fintype.prod_fiberwise [fintype α] [decidable_eq β] [fintype β] [comm_m (f : α → β) (g : α → γ) : (∏ b : β, ∏ a : {a // f a = b}, g (a : α)) = ∏ a, g a := begin - rw [← (equiv.sigma_preimage_equiv f).prod_comp, ← univ_sigma_univ, prod_sigma], + rw [← (equiv.sigma_fiber_equiv f).prod_comp, ← univ_sigma_univ, prod_sigma], refl end diff --git a/src/group_theory/coset.lean b/src/group_theory/coset.lean index 3aa01e43342bf..84d750bca52d5 100644 --- a/src/group_theory/coset.lean +++ b/src/group_theory/coset.lean @@ -383,7 +383,7 @@ def right_coset_equiv_subgroup (g : α) : right_coset ↑s g ≃ s := noncomputable def group_equiv_quotient_times_subgroup : α ≃ (α ⧸ s) × s := calc α ≃ Σ L : α ⧸ s, {x : α // (x : α ⧸ s) = L} : - (equiv.sigma_preimage_equiv quotient_group.mk).symm + (equiv.sigma_fiber_equiv quotient_group.mk).symm ... ≃ Σ L : α ⧸ s, left_coset (quotient.out' L) s : equiv.sigma_congr_right (λ L, begin diff --git a/src/group_theory/group_action/basic.lean b/src/group_theory/group_action/basic.lean index c95c5a19ebfb7..88f649e9ab8ff 100644 --- a/src/group_theory/group_action/basic.lean +++ b/src/group_theory/group_action/basic.lean @@ -243,7 +243,7 @@ as a special case."] def self_equiv_sigma_orbits' {φ : Ω → β} (hφ : right_inverse φ quotient.mk') : β ≃ Σ (ω : Ω), orbit α (φ ω) := calc β - ≃ Σ (ω : Ω), {b // quotient.mk' b = ω} : (equiv.sigma_preimage_equiv quotient.mk').symm + ≃ Σ (ω : Ω), {b // quotient.mk' b = ω} : (equiv.sigma_fiber_equiv quotient.mk').symm ... ≃ Σ (ω : Ω), orbit α (φ ω) : equiv.sigma_congr_right (λ ω, equiv.subtype_equiv_right $ λ x, by {rw [← hφ ω, quotient.eq', hφ ω], refl }) diff --git a/src/group_theory/p_group.lean b/src/group_theory/p_group.lean index 833a2c7c4658a..cd1ac4f63fb67 100644 --- a/src/group_theory/p_group.lean +++ b/src/group_theory/p_group.lean @@ -121,7 +121,7 @@ lemma card_modeq_card_fixed_points : card α ≡ card (fixed_points G α) [MOD p begin classical, calc card α = card (Σ y : quotient (orbit_rel G α), {x // quotient.mk' x = y}) : - card_congr (equiv.sigma_preimage_equiv (@quotient.mk' _ (orbit_rel G α))).symm + card_congr (equiv.sigma_fiber_equiv (@quotient.mk' _ (orbit_rel G α))).symm ... = ∑ a : quotient (orbit_rel G α), card {x // quotient.mk' x = a} : card_sigma _ ... ≡ ∑ a : fixed_points G α, 1 [MOD p] : _ ... = _ : by simp; refl, diff --git a/src/logic/equiv/basic.lean b/src/logic/equiv/basic.lean index 0f144bf0a4210..a6d9144cb6c83 100644 --- a/src/logic/equiv/basic.lean +++ b/src/logic/equiv/basic.lean @@ -778,10 +778,11 @@ def sum_equiv_sigma_bool (α β : Type u) : α ⊕ β ≃ (Σ b: bool, cond b α λ s, by cases s; refl, λ s, by rcases s with ⟨_|_, _⟩; refl⟩ -/-- `sigma_preimage_equiv f` for `f : α → β` is the natural equivalence between +/-- `sigma_fiber_equiv f` for `f : α → β` is the natural equivalence between the type of all fibres of `f` and the total space `α`. -/ +-- See also `equiv.sigma_preimage_equiv`. @[simps] -def sigma_preimage_equiv {α β : Type*} (f : α → β) : +def sigma_fiber_equiv {α β : Type*} (f : α → β) : (Σ y : β, {x // f x = y}) ≃ α := ⟨λ x, ↑x.2, λ x, ⟨f x, x, rfl⟩, λ ⟨y, x, rfl⟩, rfl, λ x, rfl⟩ @@ -1119,6 +1120,18 @@ lemma sigma_equiv_prod_sigma_congr_right : (prod_congr_right e).trans (sigma_equiv_prod α₁ β₂).symm := by { ext ⟨a, b⟩ : 1, simp } +/-- A family of equivalences between fibers gives an equivalence between domains. -/ +-- See also `equiv.of_preimage_equiv`. +@[simps] +def of_fiber_equiv {α β γ : Type*} {f : α → γ} {g : β → γ} + (e : Π c, {a // f a = c} ≃ {b // g b = c}) : + α ≃ β := +(sigma_fiber_equiv f).symm.trans $ (equiv.sigma_congr_right e).trans (sigma_fiber_equiv g) + +lemma of_fiber_equiv_map {α β γ} {f : α → γ} {g : β → γ} + (e : Π c, {a // f a = c} ≃ {b // g b = c}) (a : α) : g (of_fiber_equiv e a) = f a := +(_ : {b // g b = _}).prop + /-- A variation on `equiv.prod_congr` where the equivalence in the second component can depend on the first component. A typical example is a shear mapping, explaining the name of this declaration. -/ @@ -1409,15 +1422,15 @@ def sigma_subtype_equiv_of_subset {α : Type u} (p : α → Type v) (q : α → /-- If a predicate `p : β → Prop` is true on the range of a map `f : α → β`, then `Σ y : {y // p y}, {x // f x = y}` is equivalent to `α`. -/ -def sigma_subtype_preimage_equiv {α : Type u} {β : Type v} (f : α → β) (p : β → Prop) +def sigma_subtype_fiber_equiv {α : Type u} {β : Type v} (f : α → β) (p : β → Prop) (h : ∀ x, p (f x)) : (Σ y : subtype p, {x : α // f x = y}) ≃ α := calc _ ≃ Σ y : β, {x : α // f x = y} : sigma_subtype_equiv_of_subset _ p (λ y ⟨x, h'⟩, h' ▸ h x) - ... ≃ α : sigma_preimage_equiv f + ... ≃ α : sigma_fiber_equiv f /-- If for each `x` we have `p x ↔ q (f x)`, then `Σ y : {y // q y}, f ⁻¹' {y}` is equivalent to `{x // p x}`. -/ -def sigma_subtype_preimage_equiv_subtype {α : Type u} {β : Type v} (f : α → β) +def sigma_subtype_fiber_equiv_subtype {α : Type u} {β : Type v} (f : α → β) {p : α → Prop} {q : β → Prop} (h : ∀ x, p x ↔ q (f x)) : (Σ y : subtype q, {x : α // f x = y}) ≃ subtype p := calc (Σ y : subtype q, {x : α // f x = y}) ≃ @@ -1430,8 +1443,7 @@ calc (Σ y : subtype q, {x : α // f x = y}) ≃ assume x, exact ⟨λ ⟨hp, h'⟩, congr_arg subtype.val h', λ h', ⟨(h x).2 (h'.symm ▸ y.2), subtype.eq h'⟩⟩ end - - ... ≃ subtype p : sigma_preimage_equiv (λ x : subtype p, (⟨f x, (h x).1 x.property⟩ : subtype q)) + ... ≃ subtype p : sigma_fiber_equiv (λ x : subtype p, (⟨f x, (h x).1 x.property⟩ : subtype q)) /-- A sigma type over an `option` is equivalent to the sigma set over the original type, if the fiber is empty at none. -/ diff --git a/src/logic/equiv/set.lean b/src/logic/equiv/set.lean index 0ac322429b1f3..57301d46efd97 100644 --- a/src/logic/equiv/set.lean +++ b/src/logic/equiv/set.lean @@ -512,6 +512,23 @@ begin by_cases hi : p i; simp [hi] end +/-- `sigma_fiber_equiv f` for `f : α → β` is the natural equivalence between +the type of all preimages of points under `f` and the total space `α`. -/ +-- See also `equiv.sigma_fiber_equiv`. +@[simps] def sigma_preimage_equiv {α β} (f : α → β) : (Σ b, f ⁻¹' {b}) ≃ α := +sigma_fiber_equiv f + +/-- A family of equivalences between preimages of points gives an equivalence between domains. -/ +-- See also `equiv.of_fiber_equiv`. +@[simps] +def of_preimage_equiv {α β γ} {f : α → γ} {g : β → γ} (e : Π c, (f ⁻¹' {c}) ≃ (g ⁻¹' {c})) : + α ≃ β := +equiv.of_fiber_equiv e + +lemma of_preimage_equiv_map {α β γ} {f : α → γ} {g : β → γ} + (e : Π c, (f ⁻¹' {c}) ≃ (g ⁻¹' {c})) (a : α) : g (of_preimage_equiv e a) = f a := +equiv.of_fiber_equiv_map e a + end equiv /-- If a function is a bijection between two sets `s` and `t`, then it induces an diff --git a/src/logic/small.lean b/src/logic/small.lean index f147a7965760a..b54426e440b75 100644 --- a/src/logic/small.lean +++ b/src/logic/small.lean @@ -136,7 +136,7 @@ small_of_injective (equiv.vector_equiv_fin α n).injective instance small_list {α : Type v} [small.{u} α] : small.{u} (list α) := begin - let e : (Σ n, vector α n) ≃ list α := equiv.sigma_preimage_equiv list.length, + let e : (Σ n, vector α n) ≃ list α := equiv.sigma_fiber_equiv list.length, exact small_of_surjective e.surjective, end diff --git a/src/set_theory/cardinal/basic.lean b/src/set_theory/cardinal/basic.lean index 693c3ec06f201..d861a75fcfb37 100644 --- a/src/set_theory/cardinal/basic.lean +++ b/src/set_theory/cardinal/basic.lean @@ -636,7 +636,7 @@ theorem sum_le_sum {ι} (f g : ι → cardinal) (H : ∀ i, f i ≤ g i) : sum f lemma mk_le_mk_mul_of_mk_preimage_le {c : cardinal} (f : α → β) (hf : ∀ b : β, #(f ⁻¹' {b}) ≤ c) : #α ≤ #β * c := -by simpa only [←mk_congr (@equiv.sigma_preimage_equiv α β f), mk_sigma, ←sum_const'] +by simpa only [←mk_congr (@equiv.sigma_fiber_equiv α β f), mk_sigma, ←sum_const'] using sum_le_sum _ _ hf /-- The range of an indexed cardinal function, whose outputs live in a higher universe than the @@ -1282,7 +1282,7 @@ mk_eq_one _ (mk_congr (equiv.vector_equiv_fin α n)).trans $ by simp theorem mk_list_eq_sum_pow (α : Type u) : #(list α) = sum (λ n : ℕ, (#α) ^ℕ n) := -calc #(list α) = #(Σ n, vector α n) : mk_congr (equiv.sigma_preimage_equiv list.length).symm +calc #(list α) = #(Σ n, vector α n) : mk_congr (equiv.sigma_fiber_equiv list.length).symm ... = sum (λ n : ℕ, (#α) ^ℕ n) : by simp theorem mk_quot_le {α : Type u} {r : α → α → Prop} : #(quot r) ≤ #α := From 28568bdbcb367cb4680e3292f02893068d27aff8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 4 May 2022 20:04:20 +0000 Subject: [PATCH 119/153] feat(set_theory/game/nim): Add basic API (#13857) --- src/set_theory/game/nim.lean | 83 +++++++++++++++++++++++++++--------- 1 file changed, 62 insertions(+), 21 deletions(-) diff --git a/src/set_theory/game/nim.lean b/src/set_theory/game/nim.lean index 197ef5036e57d..ef37d3db98c37 100644 --- a/src/set_theory/game/nim.lean +++ b/src/set_theory/game/nim.lean @@ -96,28 +96,69 @@ begin exact funext (λ i, IH _ (typein_lt_self i)) end -instance nim_impartial : ∀ (O : ordinal), impartial (nim O) -| O := +lemma left_moves_nim (O : ordinal) : (nim O).left_moves = O.out.α := +by { rw nim_def, refl } +lemma right_moves_nim (O : ordinal) : (nim O).right_moves = O.out.α := +by { rw nim_def, refl } + +lemma move_left_nim_heq (O : ordinal) : (nim O).move_left == λ i : O.out.α, nim (typein (<) i) := +by { rw nim_def, refl } +lemma move_right_nim_heq (O : ordinal) : (nim O).move_right == λ i : O.out.α, nim (typein (<) i) := +by { rw nim_def, refl } + +/-- Turns an ordinal less than `O` into a left move for `nim O` and viceversa. -/ +noncomputable def to_left_moves_nim {O : ordinal} : {O' // O' < O} ≃ (nim O).left_moves := +(out_equiv_lt O).trans (equiv.cast (left_moves_nim O).symm) + +/-- Turns an ordinal less than `O` into a right move for `nim O` and viceversa. -/ +noncomputable def to_right_moves_nim {O : ordinal} : {O' // O' < O} ≃ (nim O).right_moves := +(out_equiv_lt O).trans (equiv.cast (right_moves_nim O).symm) + +@[simp] theorem to_left_moves_nim_symm_lt {O : ordinal} (i : (nim O).left_moves) : + ↑(to_left_moves_nim.symm i) < O := +(to_left_moves_nim.symm i).prop + +@[simp] theorem to_right_moves_nim_symm_lt {O : ordinal} (i : (nim O).right_moves) : + ↑(to_right_moves_nim.symm i) < O := +(to_right_moves_nim.symm i).prop + +@[simp] lemma move_left_nim' {O : ordinal.{u}} (i) : + (nim O).move_left i = nim (to_left_moves_nim.symm i).val := +(congr_heq (move_left_nim_heq O).symm (cast_heq _ i)).symm + +lemma move_left_nim {O : ordinal} (i) : + (nim O).move_left (to_left_moves_nim i) = nim i := +by simp + +@[simp] lemma move_right_nim' {O : ordinal} (i) : + (nim O).move_right i = nim (to_right_moves_nim.symm i).val := +(congr_heq (move_right_nim_heq O).symm (cast_heq _ i)).symm + +lemma move_right_nim {O : ordinal} (i) : + (nim O).move_right (to_right_moves_nim i) = nim i := +by simp + +@[simp] lemma neg_nim (O : ordinal) : -nim O = nim O := begin - rw [impartial_def, nim_def, neg_def], - split, split; - rw pgame.le_def, - all_goals { refine ⟨λ i, let hwf := ordinal.typein_lt_self i in _, - λ j, let hwf := ordinal.typein_lt_self j in _⟩ }, - { exact or.inl ⟨i, (@impartial.neg_equiv_self _ $ nim_impartial $ typein (<) i).1⟩ }, - { exact or.inr ⟨j, (@impartial.neg_equiv_self _ $ nim_impartial $ typein (<) j).1⟩ }, - { exact or.inl ⟨i, (@impartial.neg_equiv_self _ $ nim_impartial $ typein (<) i).2⟩ }, - { exact or.inr ⟨j, (@impartial.neg_equiv_self _ $ nim_impartial $ typein (<) j).2⟩ }, - { exact nim_impartial (typein (<) i) }, - { exact nim_impartial (typein (<) j) } + induction O using ordinal.induction with O IH, + rw nim_def, dsimp; congr; + funext i; + exact IH _ (ordinal.typein_lt_self i) +end + +instance nim_impartial (O : ordinal) : impartial (nim O) := +begin + induction O using ordinal.induction with O IH, + rw [impartial_def, neg_nim], + refine ⟨equiv_refl _, λ i, _, λ i, _⟩; + simpa using IH _ (typein_lt_self _) end -using_well_founded { dec_tac := tactic.assumption } -lemma exists_ordinal_move_left_eq (O : ordinal) : ∀ i, ∃ O' < O, (nim O).move_left i = nim O' := -by { rw nim_def, exact λ i, ⟨_, ⟨ordinal.typein_lt_self i, rfl⟩⟩ } +lemma exists_ordinal_move_left_eq {O : ordinal} (i) : ∃ O' < O, (nim O).move_left i = nim O' := +⟨_, typein_lt_self _, move_left_nim' i⟩ -lemma exists_move_left_eq {O : ordinal} : ∀ O' < O, ∃ i, (nim O).move_left i = nim O' := -by { rw nim_def, exact λ _ h, ⟨(ordinal.principal_seg_out h).top, by simp⟩ } +lemma exists_move_left_eq {O O' : ordinal} (h : O' < O) : ∃ i, (nim O).move_left i = nim O' := +⟨to_left_moves_nim ⟨O', h⟩, by simp⟩ @[simp] lemma zero_first_loses : (nim (0 : ordinal)).first_loses := begin @@ -254,7 +295,7 @@ begin all_goals { -- One of the piles is reduced to `k` stones, with `k < n` or `k < m`. - obtain ⟨ok, ⟨hk, hk'⟩⟩ := nim.exists_ordinal_move_left_eq _ a, + obtain ⟨ok, hk, hk'⟩ := nim.exists_ordinal_move_left_eq a, obtain ⟨k, rfl⟩ := ordinal.lt_omega.1 (lt_trans hk (ordinal.nat_lt_omega _)), replace hk := ordinal.nat_cast_lt.1 hk, @@ -288,11 +329,11 @@ begin -- Therefore, we can play the corresponding move, and by the inductive hypothesis the new state -- is `(u xor m) xor m = u` or `n xor (u xor n) = u` as required. - { obtain ⟨i, hi⟩ := nim.exists_move_left_eq _ (ordinal.nat_cast_lt.2 h), + { obtain ⟨i, hi⟩ := nim.exists_move_left_eq (ordinal.nat_cast_lt.2 h), refine ⟨(left_moves_add _ _).symm (sum.inl i), _⟩, simp only [hi, add_move_left_inl], rw [hn _ h, nat.lxor_assoc, nat.lxor_self, nat.lxor_zero] }, - { obtain ⟨i, hi⟩ := nim.exists_move_left_eq _ (ordinal.nat_cast_lt.2 h), + { obtain ⟨i, hi⟩ := nim.exists_move_left_eq (ordinal.nat_cast_lt.2 h), refine ⟨(left_moves_add _ _).symm (sum.inr i), _⟩, simp only [hi, add_move_left_inr], rw [hm _ h, nat.lxor_comm, nat.lxor_assoc, nat.lxor_self, nat.lxor_zero] } }, From eb1a566f92f1561fa476c14bd72f37cf62257dbf Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Wed, 4 May 2022 20:04:21 +0000 Subject: [PATCH 120/153] refactor(set_theory/game/ordinal): Improve API (#13878) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We change our former equivalence `o.out.α ≃ o.to_pgame.left_moves` to an equivalence `{o' // o' < o} ≃ o.to_pgame.left_moves`. This makes two proofs much simpler. We also add a simple missing lemma, `to_pgame_equiv_iff`. Co-authored-by: Scott Morrison --- src/set_theory/game/ordinal.lean | 48 +++++++++++++++----------------- 1 file changed, 23 insertions(+), 25 deletions(-) diff --git a/src/set_theory/game/ordinal.lean b/src/set_theory/game/ordinal.lean index 4c273ff985342..8d8750d01ee48 100644 --- a/src/set_theory/game/ordinal.lean +++ b/src/set_theory/game/ordinal.lean @@ -24,6 +24,8 @@ set consists of all previous ordinals. - Prove that `birthday o.to_pgame = o`. -/ +local infix ` ≈ ` := pgame.equiv + universe u namespace ordinal @@ -44,46 +46,39 @@ by rw [to_pgame, pgame.left_moves] @[simp] theorem to_pgame_right_moves (o : ordinal) : o.to_pgame.right_moves = pempty := by rw [to_pgame, pgame.right_moves] -instance : is_empty (0 : ordinal).to_pgame.left_moves := +instance : is_empty (to_pgame 0).left_moves := by { rw to_pgame_left_moves, apply_instance } instance (o : ordinal) : is_empty o.to_pgame.right_moves := by { rw to_pgame_right_moves, apply_instance } -/-- Converts a member of `o.out.α` into a move for the `pgame` corresponding to `o`, and vice versa. +/-- Converts an ordinal less than `o` into a move for the `pgame` corresponding to `o`, and vice +versa. -/ +noncomputable def to_left_moves_to_pgame {o : ordinal} : {o' // o' < o} ≃ o.to_pgame.left_moves := +(out_equiv_lt o).trans (equiv.cast (to_pgame_left_moves o).symm) -Even though these types are the same (not definitionally so), this is the preferred way to convert -between them. -/ -def to_left_moves_to_pgame {o : ordinal} : o.out.α ≃ o.to_pgame.left_moves := -equiv.cast (to_pgame_left_moves o).symm +@[simp] theorem to_left_moves_to_pgame_symm_lt {o : ordinal} (i : o.to_pgame.left_moves) : + ↑(to_left_moves_to_pgame.symm i) < o := +(to_left_moves_to_pgame.symm i).prop theorem to_pgame_move_left_heq {o : ordinal} : o.to_pgame.move_left == λ x : o.out.α, (typein (<) x).to_pgame := by { rw to_pgame, refl } -@[simp] theorem to_pgame_move_left {o : ordinal} (i : o.out.α) : - o.to_pgame.move_left (to_left_moves_to_pgame i) = (typein (<) i).to_pgame := -by { rw to_left_moves_to_pgame, exact congr_heq to_pgame_move_left_heq (cast_heq _ i) } +@[simp] theorem to_pgame_move_left' {o : ordinal} (i) : + o.to_pgame.move_left i = (to_left_moves_to_pgame.symm i).val.to_pgame := +(congr_heq to_pgame_move_left_heq.symm (cast_heq _ i)).symm + +theorem to_pgame_move_left {o : ordinal} (i) : + o.to_pgame.move_left (to_left_moves_to_pgame i) = i.val.to_pgame := +by simp theorem to_pgame_lt {a b : ordinal} (h : a < b) : a.to_pgame < b.to_pgame := -begin - convert pgame.move_left_lt (to_left_moves_to_pgame (enum (<) a _)), - { rw [to_pgame_move_left, typein_enum] }, - { rwa type_lt } -end +by { convert pgame.move_left_lt (to_left_moves_to_pgame ⟨a, h⟩), rw to_pgame_move_left } theorem to_pgame_le {a b : ordinal} (h : a ≤ b) : a.to_pgame ≤ b.to_pgame := -begin - rw pgame.le_def, - refine ⟨λ i, or.inl ⟨to_left_moves_to_pgame - (enum (<) (typein (<) (to_left_moves_to_pgame.symm i)) _), _⟩, is_empty_elim⟩, - { rw type_lt, - apply lt_of_lt_of_le _ h, - simp_rw ←type_lt a, - apply typein_lt_type }, - { rw [←to_left_moves_to_pgame.apply_symm_apply i, to_pgame_move_left], - simp } -end +pgame.le_def.2 ⟨λ i, or.inl ⟨to_left_moves_to_pgame + ⟨_, (to_left_moves_to_pgame_symm_lt i).trans_le h⟩, by simp⟩, is_empty_elim⟩ @[simp] theorem to_pgame_lt_iff {a b : ordinal} : a.to_pgame < b.to_pgame ↔ a < b := ⟨by { contrapose, rw [not_lt, pgame.not_lt], exact to_pgame_le }, to_pgame_lt⟩ @@ -91,6 +86,9 @@ end @[simp] theorem to_pgame_le_iff {a b : ordinal} : a.to_pgame ≤ b.to_pgame ↔ a ≤ b := ⟨by { contrapose, rw [not_le, pgame.not_le], exact to_pgame_lt }, to_pgame_le⟩ +@[simp] theorem to_pgame_equiv_iff {a b : ordinal} : a.to_pgame ≈ b.to_pgame ↔ a = b := +by rw [pgame.equiv, le_antisymm_iff, to_pgame_le_iff, to_pgame_le_iff] + theorem to_pgame_injective : function.injective ordinal.to_pgame := λ a b h, begin by_contra hne, From 9503f73336f35b9e1e0c1b3a37a364768187749f Mon Sep 17 00:00:00 2001 From: antoinelab01 Date: Wed, 4 May 2022 20:04:22 +0000 Subject: [PATCH 121/153] feat(linear_algebra/dual): dual of a finite free module is finite free (#13896) --- src/linear_algebra/dual.lean | 16 +++++++++++++++- 1 file changed, 15 insertions(+), 1 deletion(-) diff --git a/src/linear_algebra/dual.lean b/src/linear_algebra/dual.lean index 0fd6a4c68cc4c..2528620e623b9 100644 --- a/src/linear_algebra/dual.lean +++ b/src/linear_algebra/dual.lean @@ -6,6 +6,8 @@ Authors: Johan Commelin, Fabian Glöckle import linear_algebra.finite_dimensional import linear_algebra.projection import linear_algebra.sesquilinear_form +import ring_theory.finiteness +import linear_algebra.free_module.finite.rank /-! # Dual vector spaces @@ -277,6 +279,18 @@ linear_equiv.of_bijective (eval R M) @[simp] lemma eval_equiv_to_linear_map {ι : Type*} [fintype ι] (b : basis ι R M) : (b.eval_equiv).to_linear_map = dual.eval R M := rfl +section + +open_locale classical + +variables [finite R M] [free R M] [nontrivial R] + +instance dual_free : free R (dual R M) := free.of_basis (free.choose_basis R M).dual_basis + +instance dual_finite : finite R (dual R M) := finite.of_basis (free.choose_basis R M).dual_basis + +end + end comm_ring /-- `simp` normal form version of `total_dual_basis` -/ @@ -565,7 +579,7 @@ open finite_dimensional variables {V₁ : Type*} [add_comm_group V₁] [module K V₁] instance [H : finite_dimensional K V] : finite_dimensional K (module.dual K V) := -linear_equiv.finite_dimensional (basis.of_vector_space K V).to_dual_equiv +by apply_instance variables [finite_dimensional K V] [finite_dimensional K V₁] From 5696275aef07e5a567d328e053f5babfddb62103 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 4 May 2022 22:19:37 +0000 Subject: [PATCH 122/153] feat(data/list/big_operators): add `list.sublist.prod_le_prod'` etc (#13879) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit * add `list.forall₂.prod_le_prod'`, `list.sublist.prod_le_prod'`, and `list.sublist_forall₂.prod_le_prod'`; * add their additive versions; * upgrade `list.forall₂_same` to an `iff`. --- src/data/list/big_operators.lean | 52 +++++++++++++++++++++++--------- src/data/list/forall2.lean | 12 +++----- src/data/multiset/powerset.lean | 4 +-- src/order/well_founded_set.lean | 14 +-------- 4 files changed, 46 insertions(+), 36 deletions(-) diff --git a/src/data/list/big_operators.lean b/src/data/list/big_operators.lean index bb70195dae307..5703dd132e3a1 100644 --- a/src/data/list/big_operators.lean +++ b/src/data/list/big_operators.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Floris van Doorn, Sébastien Gouëzel, Alex J. Best -/ import algebra.group_power -import data.list.basic +import data.list.forall2 /-! # Sums and products from lists @@ -193,15 +193,46 @@ lemma _root_.commute.list_sum_left [non_unital_non_assoc_semiring R] (b : R) (l commute l.sum b := (commute.list_sum_right _ _ $ λ x hx, (h _ hx).symm).symm +@[to_additive sum_le_sum] lemma forall₂.prod_le_prod' [preorder M] + [covariant_class M M (function.swap (*)) (≤)] [covariant_class M M (*) (≤)] + {l₁ l₂ : list M} (h : forall₂ (≤) l₁ l₂) : l₁.prod ≤ l₂.prod := +begin + induction h with a b la lb hab ih ih', + { refl }, + { simpa only [prod_cons] using mul_le_mul' hab ih' } +end + +/-- If `l₁` is a sublist of `l₂` and all elements of `l₂` are greater than or equal to one, then +`l₁.prod ≤ l₂.prod`. One can prove a stronger version assuming `∀ a ∈ l₂.diff l₁, 1 ≤ a` instead +of `∀ a ∈ l₂, 1 ≤ a` but this lemma is not yet in `mathlib`. -/ +@[to_additive sum_le_sum "If `l₁` is a sublist of `l₂` and all elements of `l₂` are nonnegative, +then `l₁.sum ≤ l₂.sum`. One can prove a stronger version assuming `∀ a ∈ l₂.diff l₁, 0 ≤ a` instead +of `∀ a ∈ l₂, 0 ≤ a` but this lemma is not yet in `mathlib`."] +lemma sublist.prod_le_prod' [preorder M] [covariant_class M M (function.swap (*)) (≤)] + [covariant_class M M (*) (≤)] {l₁ l₂ : list M} (h : l₁ <+ l₂) (h₁ : ∀ a ∈ l₂, (1 : M) ≤ a) : + l₁.prod ≤ l₂.prod := +begin + induction h, { refl }, + case cons : l₁ l₂ a ih ih' + { simp only [prod_cons, forall_mem_cons] at h₁ ⊢, + exact (ih' h₁.2).trans (le_mul_of_one_le_left' h₁.1) }, + case cons2 : l₁ l₂ a ih ih' + { simp only [prod_cons, forall_mem_cons] at h₁ ⊢, + exact mul_le_mul_left' (ih' h₁.2) _ } +end + +@[to_additive sum_le_sum] lemma sublist_forall₂.prod_le_prod' [preorder M] + [covariant_class M M (function.swap (*)) (≤)] [covariant_class M M (*) (≤)] + {l₁ l₂ : list M} (h : sublist_forall₂ (≤) l₁ l₂) (h₁ : ∀ a ∈ l₂, (1 : M) ≤ a) : + l₁.prod ≤ l₂.prod := +let ⟨l, hall, hsub⟩ := sublist_forall₂_iff.1 h +in hall.prod_le_prod'.trans $ hsub.prod_le_prod' h₁ + @[to_additive sum_le_sum] lemma prod_le_prod' [preorder M] [covariant_class M M (function.swap (*)) (≤)] [covariant_class M M (*) (≤)] {l : list ι} {f g : ι → M} (h : ∀ i ∈ l, f i ≤ g i) : (l.map f).prod ≤ (l.map g).prod := -begin - induction l with i l ihl, { refl }, - rw forall_mem_cons at h, - simpa using mul_le_mul' h.1 (ihl h.2) -end +forall₂.prod_le_prod' $ by simpa @[to_additive sum_lt_sum] lemma prod_lt_prod' [preorder M] [covariant_class M M (*) (<)] [covariant_class M M (*) (≤)] @@ -401,14 +432,7 @@ le_antisymm (hl₂ ▸ single_le_prod hl₁ _ hx) (hl₁ x hx) @[to_additive] lemma prod_eq_one_iff [canonically_ordered_monoid M] (l : list M) : l.prod = 1 ↔ ∀ x ∈ l, x = (1 : M) := ⟨all_one_of_le_one_le_of_prod_eq_one (λ _ _, one_le _), -begin - induction l, - { simp }, - { intro h, - rw [prod_cons, mul_eq_one_iff], - rw forall_mem_cons at h, - exact ⟨h.1, l_ih h.2⟩ }, -end⟩ + λ h, by rw [eq_repeat.2 ⟨rfl, h⟩, prod_repeat, one_pow]⟩ /-- If all elements in a list are bounded below by `1`, then the length of the list is bounded by the sum of the elements. -/ diff --git a/src/data/list/forall2.lean b/src/data/list/forall2.lean index 20b448701ff50..335a7149238ab 100644 --- a/src/data/list/forall2.lean +++ b/src/data/list/forall2.lean @@ -40,16 +40,14 @@ lemma forall₂.flip : ∀ {a b}, forall₂ (flip r) b a → forall₂ r a b | _ _ forall₂.nil := forall₂.nil | (a :: as) (b :: bs) (forall₂.cons h₁ h₂) := forall₂.cons h₁ h₂.flip -lemma forall₂_same {r : α → α → Prop} : ∀ {l}, (∀ x∈l, r x x) → forall₂ r l l -| [] _ := forall₂.nil -| (a :: as) h := forall₂.cons - (h _ (mem_cons_self _ _)) - (forall₂_same $ λ a ha, h a $ mem_cons_of_mem _ ha) +@[simp] lemma forall₂_same {r : α → α → Prop} : ∀ {l : list α}, forall₂ r l l ↔ ∀ x ∈ l, r x x +| [] := by simp +| (a :: l) := by simp [@forall₂_same l] lemma forall₂_refl {r} [is_refl α r] (l : list α) : forall₂ r l l := -forall₂_same $ λ a h, is_refl.refl _ +forall₂_same.2 $ λ a h, refl _ -lemma forall₂_eq_eq_eq : forall₂ ((=) : α → α → Prop) = (=) := +@[simp] lemma forall₂_eq_eq_eq : forall₂ ((=) : α → α → Prop) = (=) := begin funext a b, apply propext, split, diff --git a/src/data/multiset/powerset.lean b/src/data/multiset/powerset.lean index b9a2e6a6d12db..f712941507367 100644 --- a/src/data/multiset/powerset.lean +++ b/src/data/multiset/powerset.lean @@ -125,8 +125,8 @@ theorem revzip_powerset_aux_lemma [decidable_eq α] (l : list α) begin have : forall₂ (λ (p : multiset α × multiset α) (s : multiset α), p = (s, ↑l - s)) (revzip l') ((revzip l').map prod.fst), - { rw forall₂_map_right_iff, - apply forall₂_same, rintro ⟨s, t⟩ h, + { rw [forall₂_map_right_iff, forall₂_same], + rintro ⟨s, t⟩ h, dsimp, rw [← H h, add_tsub_cancel_left] }, rw [← forall₂_eq_eq_eq, forall₂_map_right_iff], simpa end diff --git a/src/order/well_founded_set.lean b/src/order/well_founded_set.lean index d425dd2f43372..93e24066cdf74 100644 --- a/src/order/well_founded_set.lean +++ b/src/order/well_founded_set.lean @@ -684,19 +684,7 @@ begin { exact hl _ hy }, { exact hl' _ hy } }, apply ((h.partially_well_ordered_on_sublist_forall₂ (≤)).image_of_monotone_on _).mono hl, - intros l1 l2 hl1 hl2 h12, - obtain ⟨l, hll1, hll2⟩ := list.sublist_forall₂_iff.1 h12, - refine le_trans (list.rel_prod (le_refl 1) (λ a b ab c d cd, mul_le_mul' ab cd) hll1) _, - obtain ⟨l', hl'⟩ := hll2.exists_perm_append, - rw [hl'.prod_eq, list.prod_append, ← mul_one l.prod, mul_assoc, one_mul], - apply mul_le_mul_left', - have hl's := λ x hx, hl2 x (list.subset.trans (l.subset_append_right _) hl'.symm.subset hx), - clear hl', - induction l' with x1 x2 x3 x4 x5, - { refl }, - rw [list.prod_cons, ← one_mul (1 : α)], - exact mul_le_mul' (hpos x1 (hl's x1 (list.mem_cons_self x1 x2))) - (x3 (λ x hx, hl's x (list.mem_cons_of_mem _ hx))) + exact λ l1 l2 hl1 hl2 h12, h12.prod_le_prod' (λ x hx, hpos x $ hl2 x hx) end end is_pwo From 627f81b477cdd0a00631312a3cab06be7a63940d Mon Sep 17 00:00:00 2001 From: tb65536 Date: Wed, 4 May 2022 22:19:38 +0000 Subject: [PATCH 123/153] feat(group_theory/order_of_element): The index-th power lands in the subgroup (#13890) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The PR adds a lemma stating `g ^ index H ∈ H`. I had to restate `G` to avoid the fintype assumption on `G`. Co-authored-by: Scott Morrison --- src/group_theory/order_of_element.lean | 6 +++++- 1 file changed, 5 insertions(+), 1 deletion(-) diff --git a/src/group_theory/order_of_element.lean b/src/group_theory/order_of_element.lean index 3727877d18e75..32db327703a14 100644 --- a/src/group_theory/order_of_element.lean +++ b/src/group_theory/order_of_element.lean @@ -7,7 +7,7 @@ import algebra.hom.iterate import data.nat.modeq import data.set.pointwise import dynamics.periodic_pts -import group_theory.quotient_group +import group_theory.index /-! # Order of an element @@ -656,6 +656,10 @@ end let ⟨m, hm⟩ := @order_of_dvd_card_univ _ x _ _ in by simp [hm, pow_mul, pow_order_of_eq_one] +@[to_additive] lemma subgroup.pow_index_mem {G : Type*} [group G] (H : subgroup G) + [fintype (G ⧸ H)] [normal H] (g : G) : g ^ index H ∈ H := +by rw [←eq_one_iff, quotient_group.coe_pow H, index_eq_card, pow_card_eq_one] + @[to_additive] lemma pow_eq_mod_card (n : ℕ) : x ^ n = x ^ (n % fintype.card G) := by rw [pow_eq_mod_order_of, ←nat.mod_mod_of_dvd n order_of_dvd_card_univ, From f8c303eac93bd0f48dd4c9a98b8f5772af6be0c1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 4 May 2022 22:19:39 +0000 Subject: [PATCH 124/153] refactor(order/filter/pointwise): Localize instances (#13898) Localize pointwise `filter` instances into the `pointwise` locale, as is done for `set` and `finset`. --- src/order/filter/pointwise.lean | 97 ++++++++++++++++++++++++--------- 1 file changed, 72 insertions(+), 25 deletions(-) diff --git a/src/order/filter/pointwise.lean b/src/order/filter/pointwise.lean index 0bcd7c7ad2ba4..490e514eb3054 100644 --- a/src/order/filter/pointwise.lean +++ b/src/order/filter/pointwise.lean @@ -30,10 +30,17 @@ distribute over pointwise operations. For example, `t ∈ g`. * `f -ᵥ g` (`filter.has_vsub`): Scalar subtraction, filter generated by all `x -ᵥ y` where `s ∈ f` and `t ∈ g`. -* `f • g` (`filter.has_smul`): Scalar multiplication, filter generated by all `x • y` where `s ∈ f` - and `t ∈ g`. +* `f • g` (`filter.has_scalar`): Scalar multiplication, filter generated by all `x • y` where + `s ∈ f` and `t ∈ g`. * `a +ᵥ f` (`filter.has_vadd_filter`): Translation, filter of all `a +ᵥ x` where `s ∈ f`. -* `a • f` (`filter.has_smul_filter`): Scaling, filter of all `a • s` where `s ∈ f`. +* `a • f` (`filter.has_scalar_filter`): Scaling, filter of all `a • s` where `s ∈ f`. + +## Implementation notes + +We put all instances in the locale `pointwise`, so that these instances are not available by +default. Note that we do not mark them as reducible (as argued by note [reducible non-instances]) +since we expect the locale to be open whenever the instances are actually used (and making the +instances reducible changes the behavior of `simp`. ## Tags @@ -52,9 +59,12 @@ namespace filter section one variables [has_one α] {f : filter α} {s : set α} -/-- `1 : filter α` is the set of sets containing `1 : α`. -/ -@[to_additive "`0 : filter α` is the set of sets containing `0 : α`."] -instance : has_one (filter α) := ⟨principal 1⟩ +/-- `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⟩ + +localized "attribute [instance] filter.has_one filter.has_zero" in pointwise @[simp, to_additive] lemma mem_one : s ∈ (1 : filter α) ↔ (1 : α) ∈ s := one_subset @@ -84,7 +94,11 @@ end one section mul variables [has_mul α] [has_mul β] {f f₁ f₂ g g₁ g₂ h : filter α} {s t : set α} -@[to_additive] instance : has_mul (filter α) := ⟨map₂ (*)⟩ +/-- 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₂ (*)⟩ + +localized "attribute [instance] filter.has_mul filter.has_add" in pointwise @[simp, to_additive] lemma map₂_mul : map₂ (*) f g = f * g := rfl @[to_additive] lemma mem_mul_iff : s ∈ f * g ↔ ∃ t₁ t₂, t₁ ∈ f ∧ t₂ ∈ g ∧ t₁ * t₂ ⊆ s := iff.rfl @@ -109,30 +123,42 @@ map_map₂_distrib $ map_mul m end mul -@[to_additive] -instance [semigroup α] : semigroup (filter α) := +open_locale pointwise + +/-- `filter α` is a `semigroup` under pointwise operations if `α` is.-/ +@[to_additive "`filter α` is an `add_semigroup` under pointwise operations if `α` is."] +protected def semigroup [semigroup α] : semigroup (filter α) := { mul := (*), mul_assoc := λ f g h, map₂_assoc mul_assoc } -@[to_additive] -instance [comm_semigroup α] : comm_semigroup (filter α) := +/-- `filter α` is a `comm_semigroup` under pointwise operations if `α` is. -/ +@[to_additive "`filter α` is an `add_comm_semigroup` under pointwise operations if `α` is."] +protected def comm_semigroup [comm_semigroup α] : comm_semigroup (filter α) := { mul_comm := λ f g, map₂_comm mul_comm, ..filter.semigroup } -@[to_additive] -instance [mul_one_class α] : mul_one_class (filter α) := +/-- `filter α` is a `mul_one_class` under pointwise operations if `α` is. -/ +@[to_additive "`filter α` is an `add_zero_class` under pointwise operations if `α` is."] +protected def mul_one_class [mul_one_class α] : mul_one_class (filter α) := { one := 1, mul := (*), one_mul := λ f, by simp only [←pure_one, ←map₂_mul, map₂_pure_left, one_mul, map_id'], mul_one := λ f, by simp only [←pure_one, ←map₂_mul, map₂_pure_right, mul_one, map_id'] } -@[to_additive] -instance [monoid α] : monoid (filter α) := { ..filter.mul_one_class, ..filter.semigroup } +/-- `filter α` is a `monoid` under pointwise operations if `α` is. -/ +@[to_additive "`filter α` is an `add_monoid` under pointwise operations if `α` is."] +protected def monoid [monoid α] : monoid (filter α) := +{ ..filter.mul_one_class, ..filter.semigroup } -@[to_additive] -instance [comm_monoid α] : comm_monoid (filter α) := +/-- `filter α` is a `comm_monoid` under pointwise operations if `α` is. -/ +@[to_additive "`filter α` is an `add_comm_monoid` under pointwise operations if `α` is."] +protected def comm_monoid [comm_monoid α] : comm_monoid (filter α) := { ..filter.mul_one_class, ..filter.comm_semigroup } +localized "attribute [instance] filter.mul_one_class filter.add_zero_class filter.semigroup + filter.add_semigroup filter.comm_semigroup filter.add_comm_semigroup filter.monoid + filter.add_monoid filter.comm_monoid filter.add_comm_monoid" in pointwise + section map variables [mul_one_class α] [mul_one_class β] @@ -183,7 +209,9 @@ variables [has_involutive_inv α] {f : filter α} {s : set α} @[to_additive] lemma inv_mem_inv (hs : s ∈ f) : s⁻¹ ∈ f⁻¹ := by rwa [mem_inv, inv_preimage, inv_inv] -instance : has_involutive_inv (filter α) := +/-- Inversion is involutive on `filter α` if it is on `α`. -/ +@[to_additive "Negation is involutive on `filter α` if it is on `α`."] +def has_involutive_inv : has_involutive_inv (filter α) := { inv_inv := λ f, map_map.trans $ by rw [inv_involutive.comp_self, map_id], ..filter.has_inv } @@ -208,7 +236,11 @@ end group section div variables [has_div α] {f f₁ f₂ g g₁ g₂ h : filter α} {s t : set α} -@[to_additive] instance : has_div (filter α) := ⟨map₂ (/)⟩ +/-- 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₂ (/)⟩ + +localized "attribute [instance] filter.has_div filter.has_sub" in pointwise @[simp, to_additive] lemma map₂_div : map₂ (/) f g = f / g := rfl @[to_additive] lemma mem_div : s ∈ f / g ↔ ∃ t₁ t₂, t₁ ∈ f ∧ t₂ ∈ g ∧ t₁ / t₂ ⊆ s := iff.rfl @@ -233,6 +265,8 @@ le_map₂_iff end div +open_locale pointwise + section group variables [group α] [group β] {f g : filter α} {f₂ : filter β} @@ -251,16 +285,20 @@ end group `div_inv_monoid` and `has_involutive_inv` but smaller than `group` and `group_with_zero`. -/ /-- `f / g = f * g⁻¹` for all `f g : filter α` if `a / b = a * b⁻¹` for all `a b : α`. -/ -@[to_additive "`f - g = f + -g` for all `f g : filter α` if `a - b = a + -b` for all `a b : α`."] -instance div_inv_monoid [group α] : div_inv_monoid (filter α) := +@[to_additive filter.sub_neg_monoid "`f - g = f + -g` for all `f g : filter α` if `a - b = a + -b` +for all `a b : α`."] +protected def div_inv_monoid [group α] : div_inv_monoid (filter α) := { div_eq_mul_inv := λ f g, map_map₂_distrib_right div_eq_mul_inv, ..filter.monoid, ..filter.has_inv, ..filter.has_div } /-- `f / g = f * g⁻¹` for all `f g : filter α` if `a / b = a * b⁻¹` for all `a b : α`. -/ -instance div_inv_monoid' [group_with_zero α] : div_inv_monoid (filter α) := +protected def div_inv_monoid' [group_with_zero α] : div_inv_monoid (filter α) := { div_eq_mul_inv := λ f g, map_map₂_distrib_right div_eq_mul_inv, ..filter.monoid, ..filter.has_inv, ..filter.has_div } +localized "attribute [instance] filter.div_inv_monoid filter.sub_neg_monoid filter.div_inv_monoid'" + in pointwise + /-! ### Scalar addition/multiplication of filters -/ section smul @@ -298,7 +336,10 @@ section vsub variables [has_vsub α β] {f f₁ f₂ g g₁ g₂ : filter β} {h : filter α} {s t : set β} include α -instance : has_vsub (filter α) (filter β) := ⟨map₂ (-ᵥ)⟩ +/-- 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₂ (-ᵥ)⟩ + +localized "attribute [instance] filter.has_vsub" in pointwise @[simp] lemma map₂_vsub : map₂ (-ᵥ) f g = f -ᵥ g := rfl lemma mem_vsub {s : set α} : s ∈ f -ᵥ g ↔ ∃ t₁ t₂, t₁ ∈ f ∧ t₂ ∈ g ∧ t₁ -ᵥ t₂ ⊆ s := iff.rfl @@ -320,8 +361,12 @@ end vsub section smul variables [has_scalar α β] {f f₁ f₂ : filter β} {s : set β} {a : α} -@[to_additive filter.has_vadd_filter] -instance has_scalar_filter : has_scalar α (filter β) := ⟨λ a, map ((•) a)⟩ +/-- `a • f` is the map of `f` under `a •` in locale `pointwise`. -/ +@[to_additive filter.has_vadd_filter +"`a +ᵥ f` is the map of `f` under `a +ᵥ` in locale `pointwise`."] +protected def has_scalar_filter : has_scalar α (filter β) := ⟨λ a, map ((•) a)⟩ + +localized "attribute [instance] filter.has_scalar_filter filter.has_vadd_filter" in pointwise @[simp, to_additive] lemma map_smul : map (λ b, a • b) f = a • f := rfl @[to_additive] lemma mem_smul_filter : s ∈ a • f ↔ (•) a ⁻¹' s ∈ f := iff.rfl @@ -339,6 +384,8 @@ map_mono hf end smul +open_locale pointwise + @[to_additive] instance smul_comm_class_filter [has_scalar α γ] [has_scalar β γ] [smul_comm_class α β γ] : smul_comm_class α (filter β) (filter γ) := From 9b245e23e083ff6b9ed2f47a93e412219c096449 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 5 May 2022 00:25:30 +0000 Subject: [PATCH 125/153] feat(analysis/convex/integral): drop an assumption, add a version (#13920) * add `convex.set_average_mem_closure`; * drop `is_closed s` assumption in `convex.average_mem_interior_of_set`; * add `ae_eq_const_or_norm_average_lt_of_norm_le_const`, a version of `ae_eq_const_or_norm_integral_lt_of_norm_le_const` for average. --- src/analysis/convex/integral.lean | 59 ++++++++++++++++++++----------- 1 file changed, 39 insertions(+), 20 deletions(-) diff --git a/src/analysis/convex/integral.lean b/src/analysis/convex/integral.lean index cdf2628ee826d..224b5b0f9c6e4 100644 --- a/src/analysis/convex/integral.lean +++ b/src/analysis/convex/integral.lean @@ -105,6 +105,15 @@ begin rwa [ne.def, restrict_eq_zero] end +/-- If `μ` is a non-zero finite measure on `α`, `s` is a convex set in `E`, and `f` is an integrable +function sending `μ`-a.e. points to `s`, then the average value of `f` belongs to `closure s`: +`⨍ x, f x ∂μ ∈ s`. See also `convex.center_mass_mem` for a finite sum version of this lemma. -/ +lemma convex.set_average_mem_closure {t : set α} {s : set E} (hs : convex ℝ s) + (h0 : μ t ≠ 0) (ht : μ t ≠ ∞) {f : α → E} (hfs : ∀ᵐ x ∂μ.restrict t, f x ∈ s) + (hfi : integrable_on f t μ) : + ⨍ x in t, f x ∂μ ∈ closure s := +hs.closure.set_average_mem is_closed_closure h0 ht (hfs.mono $ λ x hx, subset_closure hx) hfi + lemma convex_on.average_mem_epigraph [is_finite_measure μ] {s : set E} {g : E → ℝ} (hg : convex_on ℝ s g) (hgc : continuous_on g s) (hsc : is_closed s) (hμ : μ ≠ 0) {f : α → E} (hfs : ∀ᵐ x ∂μ, f x ∈ s) (hfi : integrable f μ) (hgi : integrable (g ∘ f) μ) : @@ -244,11 +253,11 @@ begin rw [this, measure_smul_set_average _ (measure_ne_top μ _)] end -/-- If an integrable function `f : α → E` takes values in a convex closed set `s` and for some set -`t` of positive measure, the average value of `f` over `t` belongs to the interior of `s`, then the -average of `f` over the whole space belongs to the interior of `s`. -/ +/-- If an integrable function `f : α → E` takes values in a convex set `s` and for some set `t` of +positive measure, the average value of `f` over `t` belongs to the interior of `s`, then the average +of `f` over the whole space belongs to the interior of `s`. -/ lemma convex.average_mem_interior_of_set [is_finite_measure μ] {t : set α} {s : set E} - (hs : convex ℝ s) (hsc : is_closed s) (h0 : μ t ≠ 0) {f : α → E} (hfs : ∀ᵐ x ∂μ, f x ∈ s) + (hs : convex ℝ s) (h0 : μ t ≠ 0) {f : α → E} (hfs : ∀ᵐ x ∂μ, f x ∈ s) (hfi : integrable f μ) (ht : ⨍ x in t, f x ∂μ ∈ interior s) : ⨍ x, f x ∂μ ∈ interior s := begin @@ -256,8 +265,8 @@ begin by_cases h0' : μ (to_measurable μ t)ᶜ = 0, { rw ← ae_eq_univ at h0', rwa [restrict_congr_set h0', restrict_univ] at ht }, - exact hs.open_segment_interior_self_subset_interior ht - (hs.set_average_mem hsc h0' (measure_ne_top _ _) (ae_restrict_of_ae hfs) hfi.integrable_on) + exact hs.open_segment_interior_closure_subset_interior ht + (hs.set_average_mem_closure h0' (measure_ne_top _ _) (ae_restrict_of_ae hfs) hfi.integrable_on) (average_mem_open_segment_compl_self (measurable_set_to_measurable μ t).null_measurable_set h0 h0' hfi) end @@ -316,27 +325,37 @@ by simpa only [pi.neg_apply, average_neg, neg_lt_neg_iff] using hg.neg.ae_eq_const_or_map_average_lt hgc.neg hsc hfs hfi hgi.neg /-- If `E` is a strictly normed space and `f : α → E` is a function such that `∥f x∥ ≤ C` a.e., then -either either this function is a.e. equal to its average value, or the norm of its integral is -strictly less than `(μ univ).to_real * C`. -/ -lemma ae_eq_const_or_norm_integral_lt_of_norm_le_const [strict_convex_space ℝ E] - [is_finite_measure μ] {f : α → E} {C : ℝ} (h_le : ∀ᵐ x ∂μ, ∥f x∥ ≤ C) : - (f =ᵐ[μ] const α ⨍ x, f x ∂μ) ∨ ∥∫ x, f x ∂μ∥ < (μ univ).to_real * C := +either this function is a.e. equal to its average value, or the norm of its average value is +strictly less than `C`. -/ +lemma ae_eq_const_or_norm_average_lt_of_norm_le_const [strict_convex_space ℝ E] + {f : α → E} {C : ℝ} (h_le : ∀ᵐ x ∂μ, ∥f x∥ ≤ C) : + (f =ᵐ[μ] const α ⨍ x, f x ∂μ) ∨ ∥⨍ x, f x ∂μ∥ < C := begin cases le_or_lt C 0 with hC0 hC0, { have : f =ᵐ[μ] 0, from h_le.mono (λ x hx, norm_le_zero_iff.1 (hx.trans hC0)), simp only [average_congr this, pi.zero_apply, average_zero], exact or.inl this }, - cases eq_or_ne μ 0 with hμ hμ, - { rw hμ, exact or.inl rfl }, by_cases hfi : integrable f μ, swap, - { right, - simpa [integral_undef hfi, hC0, measure_lt_top, ennreal.to_real_pos_iff, pos_iff_ne_zero] - using hμ }, + by simp [average_def', integral_undef hfi, hC0, ennreal.to_real_pos_iff], + cases (le_top : μ univ ≤ ∞).eq_or_lt with hμt hμt, { simp [average_def', hμt, hC0] }, + haveI : is_finite_measure μ := ⟨hμt⟩, replace h_le : ∀ᵐ x ∂μ, f x ∈ closed_ball (0 : E) C, by simpa only [mem_closed_ball_zero_iff], - have hμ' : 0 < (μ univ).to_real, - from ennreal.to_real_pos (mt measure_univ_eq_zero.1 hμ) (measure_ne_top _ _), - simpa only [interior_closed_ball _ hC0.ne', mem_ball_zero_iff, average_def', norm_smul, - real.norm_eq_abs, abs_inv, abs_of_pos hμ', ← div_eq_inv_mul, div_lt_iff' hμ'] + simpa only [interior_closed_ball _ hC0.ne', mem_ball_zero_iff] using (strict_convex_closed_ball ℝ (0 : E) C).ae_eq_const_or_average_mem_interior is_closed_ball h_le hfi end + +/-- If `E` is a strictly normed space and `f : α → E` is a function such that `∥f x∥ ≤ C` a.e., then +either this function is a.e. equal to its average value, or the norm of its integral is strictly +less than `(μ univ).to_real * C`. -/ +lemma ae_eq_const_or_norm_integral_lt_of_norm_le_const [strict_convex_space ℝ E] + [is_finite_measure μ] {f : α → E} {C : ℝ} (h_le : ∀ᵐ x ∂μ, ∥f x∥ ≤ C) : + (f =ᵐ[μ] const α ⨍ x, f x ∂μ) ∨ ∥∫ x, f x ∂μ∥ < (μ univ).to_real * C := +begin + cases eq_or_ne μ 0 with h₀ h₀, { left, simp [h₀] }, + have hμ : 0 < (μ univ).to_real, + by simp [ennreal.to_real_pos_iff, pos_iff_ne_zero, h₀, measure_lt_top], + refine (ae_eq_const_or_norm_average_lt_of_norm_le_const h_le).imp_right (λ H, _), + rwa [average_def', norm_smul, norm_inv, real.norm_eq_abs, abs_of_pos hμ, + ← div_eq_inv_mul, div_lt_iff' hμ] at H +end From 3650936965ead994c067cbce455aa5eb60cd0e55 Mon Sep 17 00:00:00 2001 From: antoinelab01 Date: Thu, 5 May 2022 00:25:31 +0000 Subject: [PATCH 126/153] feat(representation_theory/Action): lemma about isomorphisms in `Action G V` (#13951) --- src/representation_theory/Action.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/src/representation_theory/Action.lean b/src/representation_theory/Action.lean index d5b408ec7cb18..270fd0ad25b90 100644 --- a/src/representation_theory/Action.lean +++ b/src/representation_theory/Action.lean @@ -14,6 +14,7 @@ import category_theory.monoidal.linear import category_theory.monoidal.braided import category_theory.abelian.functor_category import category_theory.abelian.transfer +import category_theory.conj import category_theory.linear.functor_category /-! @@ -246,6 +247,10 @@ preserves_colimits_of_nat_iso end forget +lemma iso.conj_ρ {M N : Action V G} (f : M ≅ N) (g : G) : + N.ρ g = (((forget V G).map_iso f).conj (M.ρ g)) := +by { rw [iso.conj_apply, iso.eq_inv_comp], simp [f.hom.comm'] } + section has_zero_morphisms variables [has_zero_morphisms V] From 1e1893566a84cbbc43b4619c4be5b441a13d5736 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Thu, 5 May 2022 00:25:32 +0000 Subject: [PATCH 127/153] docs(algebra/ring/opposite): fix docstring for `ring_hom.from_opposite` (#13957) --- src/algebra/ring/opposite.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/algebra/ring/opposite.lean b/src/algebra/ring/opposite.lean index 8497108915589..a833861e4dd15 100644 --- a/src/algebra/ring/opposite.lean +++ b/src/algebra/ring/opposite.lean @@ -176,8 +176,8 @@ def ring_hom.to_opposite {R S : Type*} [semiring R] [semiring S] (f : R →+* S) .. ((op_add_equiv : S ≃+ Sᵐᵒᵖ).to_add_monoid_hom.comp ↑f : R →+ Sᵐᵒᵖ), .. f.to_monoid_hom.to_opposite hf } -/-- A monoid homomorphism `f : R →* S` such that `f x` commutes with `f y` for all `x, y` defines -a monoid homomorphism from `Rᵐᵒᵖ`. -/ +/-- A ring homomorphism `f : R →+* S` such that `f x` commutes with `f y` for all `x, y` defines +a ring homomorphism from `Rᵐᵒᵖ`. -/ @[simps {fully_applied := ff}] def ring_hom.from_opposite {R S : Type*} [semiring R] [semiring S] (f : R →+* S) (hf : ∀ x y, commute (f x) (f y)) : Rᵐᵒᵖ →+* S := From 50fd3d6adfbac00997babac42d704d7e5415a7b5 Mon Sep 17 00:00:00 2001 From: Bolton Bailey Date: Thu, 5 May 2022 02:21:58 +0000 Subject: [PATCH 128/153] feat(analysis/special_functions/log/monotone): add lemmas (#13848) Adds a few lemmas regarding tonality of `log x / x ^ a`, and puts them in a new file, along with previous results. --- src/analysis/special_functions/arsinh.lean | 2 +- .../special_functions/complex/log.lean | 2 +- .../{logb.lean => log/base.lean} | 2 +- .../{log.lean => log/basic.lean} | 16 ---- .../{log_deriv.lean => log/deriv.lean} | 2 +- .../special_functions/log/monotone.lean | 93 +++++++++++++++++++ src/analysis/special_functions/pow_deriv.lean | 2 +- src/data/complex/exponential_bounds.lean | 2 +- src/number_theory/von_mangoldt.lean | 2 +- test/differentiable.lean | 2 +- 10 files changed, 101 insertions(+), 24 deletions(-) rename src/analysis/special_functions/{logb.lean => log/base.lean} (99%) rename src/analysis/special_functions/{log.lean => log/basic.lean} (92%) rename src/analysis/special_functions/{log_deriv.lean => log/deriv.lean} (99%) create mode 100644 src/analysis/special_functions/log/monotone.lean diff --git a/src/analysis/special_functions/arsinh.lean b/src/analysis/special_functions/arsinh.lean index f77d3cc0df16f..de33a1a039dc7 100644 --- a/src/analysis/special_functions/arsinh.lean +++ b/src/analysis/special_functions/arsinh.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: James Arthur, Chris Hughes, Shing Tak Lam -/ import analysis.special_functions.trigonometric.deriv -import analysis.special_functions.log +import analysis.special_functions.log.basic /-! # Inverse of the sinh function diff --git a/src/analysis/special_functions/complex/log.lean b/src/analysis/special_functions/complex/log.lean index 50ae4eb769cf1..ca30aaa89c548 100644 --- a/src/analysis/special_functions/complex/log.lean +++ b/src/analysis/special_functions/complex/log.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Abhimanyu Pallavi Sudhir, Jean Lo, Calle Sönne, Benjamin Davidson -/ import analysis.special_functions.complex.arg -import analysis.special_functions.log +import analysis.special_functions.log.basic /-! # The complex `log` function diff --git a/src/analysis/special_functions/logb.lean b/src/analysis/special_functions/log/base.lean similarity index 99% rename from src/analysis/special_functions/logb.lean rename to src/analysis/special_functions/log/base.lean index 50a37e1bc7454..498abcd8cd450 100644 --- a/src/analysis/special_functions/logb.lean +++ b/src/analysis/special_functions/log/base.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Bolton Bailey. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Bolton Bailey, Chris Hughes, Abhimanyu Pallavi Sudhir, Jean Lo, Calle Sönne -/ -import analysis.special_functions.log +import analysis.special_functions.log.basic import analysis.special_functions.pow /-! diff --git a/src/analysis/special_functions/log.lean b/src/analysis/special_functions/log/basic.lean similarity index 92% rename from src/analysis/special_functions/log.lean rename to src/analysis/special_functions/log/basic.lean index 0e8b896ccd9a6..91f8e7b559953 100644 --- a/src/analysis/special_functions/log.lean +++ b/src/analysis/special_functions/log/basic.lean @@ -196,22 +196,6 @@ begin rw exp_log hx, end -lemma log_div_self_antitone_on : antitone_on (λ x : ℝ, log x / x) {x | exp 1 ≤ x} := -begin - simp only [antitone_on, mem_set_of_eq], - intros x hex y hey hxy, - have x_pos : 0 < x := (exp_pos 1).trans_le hex, - have y_pos : 0 < y := (exp_pos 1).trans_le hey, - have hlogx : 1 ≤ log x := by rwa le_log_iff_exp_le x_pos, - have hyx : 0 ≤ y / x - 1 := by rwa [le_sub_iff_add_le, le_div_iff x_pos, zero_add, one_mul], - rw [div_le_iff y_pos, ←sub_le_sub_iff_right (log x)], - calc - log y - log x = log (y / x) : by rw [log_div (y_pos.ne') (x_pos.ne')] - ... ≤ (y / x) - 1 : log_le_sub_one_of_pos (div_pos y_pos x_pos) - ... ≤ log x * (y / x - 1) : le_mul_of_one_le_left hyx hlogx - ... = log x / x * y - log x : by ring, -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/log_deriv.lean b/src/analysis/special_functions/log/deriv.lean similarity index 99% rename from src/analysis/special_functions/log_deriv.lean rename to src/analysis/special_functions/log/deriv.lean index 957dd9f638a54..af95ec6e16085 100644 --- a/src/analysis/special_functions/log_deriv.lean +++ b/src/analysis/special_functions/log/deriv.lean @@ -3,7 +3,7 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Abhimanyu Pallavi Sudhir, Jean Lo, Calle Sönne -/ -import analysis.special_functions.log +import analysis.special_functions.log.basic import analysis.special_functions.exp_deriv /-! diff --git a/src/analysis/special_functions/log/monotone.lean b/src/analysis/special_functions/log/monotone.lean new file mode 100644 index 0000000000000..665632df01b50 --- /dev/null +++ b/src/analysis/special_functions/log/monotone.lean @@ -0,0 +1,93 @@ +/- +Copyright (c) 2021 Bolton Bailey. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Bolton Bailey +-/ +import analysis.special_functions.log.basic +import analysis.special_functions.pow + +/-! +# Logarithm Tonality + +In this file we describe the tonality of the logarithm function when multiplied by functions of the +form `x ^ a`. + +## Tags + +logarithm, tonality +-/ + +open set filter function +open_locale topological_space +noncomputable theory + +namespace real +variables {x y : ℝ} + + +lemma log_mul_self_monotone_on : monotone_on (λ x : ℝ, log x * x) {x | 1 ≤ x} := +begin + -- TODO: can be strengthened to exp (-1) ≤ x + simp only [monotone_on, mem_set_of_eq], + intros x hex y hey hxy, + have x_pos : 0 < x := lt_of_lt_of_le zero_lt_one hex, + have y_pos : 0 < y := lt_of_lt_of_le zero_lt_one hey, + refine mul_le_mul ((log_le_log x_pos y_pos).mpr hxy) hxy (le_of_lt x_pos) _, + rwa [le_log_iff_exp_le y_pos, real.exp_zero], +end + +lemma log_div_self_antitone_on : antitone_on (λ x : ℝ, log x / x) {x | exp 1 ≤ x} := +begin + simp only [antitone_on, mem_set_of_eq], + intros x hex y hey hxy, + have x_pos : 0 < x := (exp_pos 1).trans_le hex, + have y_pos : 0 < y := (exp_pos 1).trans_le hey, + have hlogx : 1 ≤ log x := by rwa le_log_iff_exp_le x_pos, + have hyx : 0 ≤ y / x - 1 := by rwa [le_sub_iff_add_le, le_div_iff x_pos, zero_add, one_mul], + rw [div_le_iff y_pos, ←sub_le_sub_iff_right (log x)], + calc + log y - log x = log (y / x) : by rw [log_div (y_pos.ne') (x_pos.ne')] + ... ≤ (y / x) - 1 : log_le_sub_one_of_pos (div_pos y_pos x_pos) + ... ≤ log x * (y / x - 1) : le_mul_of_one_le_left hyx hlogx + ... = log x / x * y - log x : by ring, +end + +lemma log_div_self_rpow_antitone_on {a : ℝ} (ha : 0 < a) : + antitone_on (λ x : ℝ, log x / x ^ a) {x | exp (1 / a) ≤ x} := +begin + simp only [antitone_on, mem_set_of_eq], + intros x hex y hey hxy, + have x_pos : 0 < x := lt_of_lt_of_le (exp_pos (1 / a)) hex, + have y_pos : 0 < y := by linarith, + have x_nonneg : 0 ≤ x := le_trans (le_of_lt (exp_pos (1 / a))) hex, + have y_nonneg : 0 ≤ y := by linarith, + nth_rewrite 0 ←rpow_one y, + nth_rewrite 0 ←rpow_one x, + rw [←div_self (ne_of_lt ha).symm, div_eq_mul_one_div a a, rpow_mul y_nonneg, rpow_mul x_nonneg, + log_rpow (rpow_pos_of_pos y_pos a), log_rpow (rpow_pos_of_pos x_pos a), mul_div_assoc, + mul_div_assoc, mul_le_mul_left (one_div_pos.mpr ha)], + { refine log_div_self_antitone_on _ _ _, + { simp only [set.mem_set_of_eq], + convert rpow_le_rpow _ hex (le_of_lt ha), + rw ←exp_mul, + simp only [real.exp_eq_exp], + field_simp [(ne_of_lt ha).symm], + exact le_of_lt (exp_pos (1 / a)), }, + { simp only [set.mem_set_of_eq], + convert rpow_le_rpow _ (trans hex hxy) (le_of_lt ha), + rw ←exp_mul, + simp only [real.exp_eq_exp], + field_simp [(ne_of_lt ha).symm], + exact le_of_lt (exp_pos (1 / a)), }, + exact rpow_le_rpow x_nonneg hxy (le_of_lt ha), }, +end + +lemma log_div_sqrt_antitone_on : + antitone_on (λ x : ℝ, log x / sqrt x) {x | exp 2 ≤ x} := +begin + simp_rw sqrt_eq_rpow, + convert @log_div_self_rpow_antitone_on (1 / 2) (by norm_num), + norm_num, +end + +end real diff --git a/src/analysis/special_functions/pow_deriv.lean b/src/analysis/special_functions/pow_deriv.lean index 86d66bc87466a..f5c538200f18b 100644 --- a/src/analysis/special_functions/pow_deriv.lean +++ b/src/analysis/special_functions/pow_deriv.lean @@ -7,7 +7,7 @@ Authors: Chris Hughes, Abhimanyu Pallavi Sudhir, Jean Lo, Calle Sönne, Sébasti import analysis.special_functions.pow import analysis.special_functions.complex.log_deriv import analysis.calculus.extend_deriv -import analysis.special_functions.log_deriv +import analysis.special_functions.log.deriv import analysis.special_functions.trigonometric.deriv /-! diff --git a/src/data/complex/exponential_bounds.lean b/src/data/complex/exponential_bounds.lean index 674a605153f41..28bffde9acb89 100644 --- a/src/data/complex/exponential_bounds.lean +++ b/src/data/complex/exponential_bounds.lean @@ -5,7 +5,7 @@ Authors: Mario Carneiro, Joseph Myers -/ import data.complex.exponential -import analysis.special_functions.log_deriv +import analysis.special_functions.log.deriv /-! # Bounds on specific values of the exponential diff --git a/src/number_theory/von_mangoldt.lean b/src/number_theory/von_mangoldt.lean index 6f96ac672b62f..999937c3d0cb5 100644 --- a/src/number_theory/von_mangoldt.lean +++ b/src/number_theory/von_mangoldt.lean @@ -6,7 +6,7 @@ Authors: Bhavik Mehta import algebra.is_prime_pow import number_theory.arithmetic_function -import analysis.special_functions.log +import analysis.special_functions.log.basic /-! # The von Mangoldt Function diff --git a/test/differentiable.lean b/test/differentiable.lean index d9c6dec1d79fe..567543cfa0078 100644 --- a/test/differentiable.lean +++ b/test/differentiable.lean @@ -1,5 +1,5 @@ import analysis.special_functions.trigonometric.deriv -import analysis.special_functions.log_deriv +import analysis.special_functions.log.deriv namespace real From 50781196cfb96cc78f4fbe3675f80d442725826b Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 5 May 2022 02:21:58 +0000 Subject: [PATCH 129/153] feat(data/matrix/block): add `matrix.block_diag` and `matrix.block_diag'` (#13918) `matrix.block_diag` is to `matrix.block_diagonal` as `matrix.diag` is to `matrix.diagonal`. As well as the basic arithmetic lemmas and bundling, this also adds continuity lemmas. These definitions are primarily an auxiliary construction to prove `matrix.tsum_block_diagonal`, and `matrix.tsum_block_diagonal'`, which are really the main goal of this PR. --- src/data/matrix/block.lean | 155 ++++++++++++++++++++++++++++- src/topology/instances/matrix.lean | 71 ++++++++++++- 2 files changed, 223 insertions(+), 3 deletions(-) diff --git a/src/data/matrix/block.lean b/src/data/matrix/block.lean index cf1c1a74e5137..108c572f023d6 100644 --- a/src/data/matrix/block.lean +++ b/src/data/matrix/block.lean @@ -15,8 +15,10 @@ import data.matrix.basic extract each of the four blocks from `matrix.from_blocks`. * `matrix.block_diagonal`: block diagonal of equally sized blocks. On square blocks, this is a ring homomorphisms, `matrix.block_diagonal_ring_hom`. +* `matrix.block_diag`: extract the blocks from the diagonal of a block diagonal matrix. * `matrix.block_diagonal'`: block diagonal of unequally sized blocks. On square blocks, this is a ring homomorphisms, `matrix.block_diagonal'_ring_hom`. +* `matrix.block_diag'`: extract the blocks from the diagonal of a block diagonal matrix. -/ variables {l m n o p q : Type*} {m' n' p' : o → Type*} @@ -334,12 +336,87 @@ by { ext, simp only [block_diagonal_apply, pi.smul_apply], split_ifs; simp } end block_diagonal +section block_diag + + +/-- Extract a block from the diagonal of a block diagonal matrix. + +This is the block form of `matrix.diag`, and the left-inverse of `matrix.block_diagonal`. -/ +def block_diag (M : matrix (m × o) (n × o) α) (k : o) : matrix m n α +| i j := M (i, k) (j, k) + +lemma block_diag_map (M : matrix (m × o) (n × o) α) (f : α → β) : + block_diag (M.map f) = λ k, (block_diag M k).map f := +rfl + +@[simp] lemma block_diag_transpose (M : matrix (m × o) (n × o) α) (k : o) : + block_diag Mᵀ k = (block_diag M k)ᵀ := +ext $ λ i j, rfl + +@[simp] lemma block_diag_conj_transpose + {α : Type*} [add_monoid α] [star_add_monoid α] (M : matrix (m × o) (n × o) α) (k : o) : + block_diag Mᴴ k = (block_diag M k)ᴴ := +ext $ λ i j, rfl + +section has_zero +variables [has_zero α] [has_zero β] + +@[simp] lemma block_diag_zero : + block_diag (0 : matrix (m × o) (n × o) α) = 0 := +rfl + +@[simp] lemma block_diag_diagonal [decidable_eq o] [decidable_eq m] (d : (m × o) → α) (k : o) : + block_diag (diagonal d) k = diagonal (λ i, d (i, k)) := +ext $ λ i j, begin + obtain rfl | hij := decidable.eq_or_ne i j, + { rw [block_diag, diagonal_apply_eq, diagonal_apply_eq] }, + { rw [block_diag, diagonal_apply_ne _ hij, diagonal_apply_ne _ (mt _ hij)], + exact prod.fst_eq_iff.mpr }, +end + +@[simp] lemma block_diag_block_diagonal [decidable_eq o] (M : o → matrix m n α) : + block_diag (block_diagonal M) = M := +funext $ λ k, ext $ λ i j, block_diagonal_apply_eq _ _ _ _ + +@[simp] lemma block_diag_one [decidable_eq o] [decidable_eq m] [has_one α] : + block_diag (1 : matrix (m × o) (m × o) α) = 1 := +funext $ block_diag_diagonal _ + +end has_zero + +@[simp] lemma block_diag_add [add_zero_class α] (M N : matrix (m × o) (n × o) α) : + block_diag (M + N) = block_diag M + block_diag N := +rfl + +section +variables (o m n α) +/-- `matrix.block_diag` as an `add_monoid_hom`. -/ +@[simps] def block_diag_add_monoid_hom [add_zero_class α] : + matrix (m × o) (n × o) α →+ (o → matrix m n α) := +{ to_fun := block_diag, + map_zero' := block_diag_zero, + map_add' := block_diag_add } +end + +@[simp] lemma block_diag_neg [add_group α] (M : matrix (m × o) (n × o) α) : + block_diag (-M) = - block_diag M := +map_neg (block_diag_add_monoid_hom m n o α) M + +@[simp] lemma block_diag_sub [add_group α] (M N : matrix (m × o) (n × o) α) : + block_diag (M - N) = block_diag M - block_diag N := +map_sub (block_diag_add_monoid_hom m n o α) M N + +@[simp] lemma block_diag_smul {R : Type*} [monoid R] [add_monoid α] [distrib_mul_action R α] + (x : R) (M : matrix (m × o) (n × o) α) : block_diag (x • M) = x • block_diag M := +rfl + +end block_diag + section block_diagonal' variables [decidable_eq o] section has_zero - variables [has_zero α] [has_zero β] /-- `matrix.block_diagonal' M` turns `M : Π i, matrix (m i) (n i) α` into a @@ -481,4 +558,80 @@ by { ext, simp only [block_diagonal'_apply, pi.smul_apply], split_ifs; simp } end block_diagonal' +section block_diag' + +/-- Extract a block from the diagonal of a block diagonal matrix. + +This is the block form of `matrix.diag`, and the left-inverse of `matrix.block_diagonal'`. -/ +def block_diag' (M : matrix (Σ i, m' i) (Σ i, n' i) α) (k : o) : matrix (m' k) (n' k) α +| i j := M ⟨k, i⟩ ⟨k, j⟩ + +lemma block_diag'_map (M : matrix (Σ i, m' i) (Σ i, n' i) α) (f : α → β) : + block_diag' (M.map f) = λ k, (block_diag' M k).map f := +rfl + +@[simp] lemma block_diag'_transpose (M : matrix (Σ i, m' i) (Σ i, n' i) α) (k : o) : + block_diag' Mᵀ k = (block_diag' M k)ᵀ := +ext $ λ i j, rfl + +@[simp] lemma block_diag'_conj_transpose + {α : Type*} [add_monoid α] [star_add_monoid α] (M : matrix (Σ i, m' i) (Σ i, n' i) α) (k : o) : + block_diag' Mᴴ k = (block_diag' M k)ᴴ := +ext $ λ i j, rfl + +section has_zero +variables [has_zero α] [has_zero β] + +@[simp] lemma block_diag'_zero : + block_diag' (0 : matrix (Σ i, m' i) (Σ i, n' i) α) = 0 := +rfl + +@[simp] lemma block_diag'_diagonal [decidable_eq o] [Π i, decidable_eq (m' i)] + (d : (Σ i, m' i) → α) (k : o) : + block_diag' (diagonal d) k = diagonal (λ i, d ⟨k, i⟩) := +ext $ λ i j, begin + obtain rfl | hij := decidable.eq_or_ne i j, + { rw [block_diag', diagonal_apply_eq, diagonal_apply_eq] }, + { rw [block_diag', diagonal_apply_ne _ hij, diagonal_apply_ne _ (mt (λ h, _) hij)], + cases h, refl }, +end + +@[simp] lemma block_diag'_block_diagonal' [decidable_eq o] (M : Π i, matrix (m' i) (n' i) α) : + block_diag' (block_diagonal' M) = M := +funext $ λ k, ext $ λ i j, block_diagonal'_apply_eq _ _ _ _ + +@[simp] lemma block_diag'_one [decidable_eq o] [Π i, decidable_eq (m' i)] [has_one α] : + block_diag' (1 : matrix (Σ i, m' i) (Σ i, m' i) α) = 1 := +funext $ block_diag'_diagonal _ + +end has_zero + +@[simp] lemma block_diag'_add [add_zero_class α] (M N : matrix (Σ i, m' i) (Σ i, n' i) α) : + block_diag' (M + N) = block_diag' M + block_diag' N := +rfl + +section +variables (m' n' α) +/-- `matrix.block_diag'` as an `add_monoid_hom`. -/ +@[simps] def block_diag'_add_monoid_hom [add_zero_class α] : + matrix (Σ i, m' i) (Σ i, n' i) α →+ Π i, matrix (m' i) (n' i) α := +{ to_fun := block_diag', + map_zero' := block_diag'_zero, + map_add' := block_diag'_add } +end + +@[simp] lemma block_diag'_neg [add_group α] (M : matrix (Σ i, m' i) (Σ i, n' i) α) : + block_diag' (-M) = - block_diag' M := +map_neg (block_diag'_add_monoid_hom m' n' α) M + +@[simp] lemma block_diag'_sub [add_group α] (M N : matrix (Σ i, m' i) (Σ i, n' i) α) : + block_diag' (M - N) = block_diag' M - block_diag' N := +map_sub (block_diag'_add_monoid_hom m' n' α) M N + +@[simp] lemma block_diag'_smul {R : Type*} [monoid R] [add_monoid α] [distrib_mul_action R α] + (x : R) (M : matrix (Σ i, m' i) (Σ i, n' i) α) : block_diag' (x • M) = x • block_diag' M := +rfl + +end block_diag' + end matrix diff --git a/src/topology/instances/matrix.lean b/src/topology/instances/matrix.lean index 319ae93cc7f2a..e6ecbaf2d458a 100644 --- a/src/topology/instances/matrix.lean +++ b/src/topology/instances/matrix.lean @@ -25,6 +25,8 @@ This file is a place to collect topological results about matrices. * Infinite sums * `matrix.transpose_tsum`: transpose commutes with infinite sums * `matrix.diagonal_tsum`: diagonal commutes with infinite sums + * `matrix.block_diagonal_tsum`: block diagonal commutes with infinite sums + * `matrix.block_diagonal'_tsum`: non-uniform block diagonal commutes with infinite sums -/ open matrix @@ -224,6 +226,11 @@ lemma continuous.matrix_block_diagonal [has_zero R] [decidable_eq p] {A : X → continuous_matrix $ λ ⟨i₁, i₂⟩ ⟨j₁, j₂⟩, (((continuous_apply i₂).comp hA).matrix_elem i₁ j₁).if_const _ continuous_zero +@[continuity] +lemma continuous.matrix_block_diag {A : X → matrix (m × p) (n × p) R} (hA : continuous A) : + continuous (λ x, block_diag (A x)) := +continuous_pi $ λ i, continuous_matrix $ λ j k, hA.matrix_elem _ _ + @[continuity] lemma continuous.matrix_block_diagonal' [has_zero R] [decidable_eq l] {A : X → Π i, matrix (m' i) (n' i) R} (hA : continuous A) : @@ -236,6 +243,11 @@ continuous_matrix $ λ ⟨i₁, i₂⟩ ⟨j₁, j₂⟩, begin { exact continuous_const }, end +@[continuity] +lemma continuous.matrix_block_diag' {A : X → matrix (Σ i, m' i) (Σ i, n' i) R} (hA : continuous A) : + continuous (λ x, block_diag' (A x)) := +continuous_pi $ λ i, continuous_matrix $ λ j k, hA.matrix_elem _ _ + end block_matrices end continuity @@ -335,17 +347,72 @@ lemma summable.matrix_block_diagonal [decidable_eq p] {f : X → p → matrix m summable (λ x, block_diagonal (f x)) := hf.has_sum.matrix_block_diagonal.summable +lemma summable_matrix_block_diagonal [decidable_eq p] {f : X → p → matrix m n R} : + summable (λ x, block_diagonal (f x)) ↔ summable f := +(summable.map_iff_of_left_inverse + (matrix.block_diagonal_add_monoid_hom m n p R) (matrix.block_diag_add_monoid_hom m n p R) + (by exact continuous.matrix_block_diagonal continuous_id) + (by exact continuous.matrix_block_diag continuous_id) + (λ A, block_diag_block_diagonal A) : _) + +lemma matrix.block_diagonal_tsum [decidable_eq p] [t2_space R] {f : X → p → matrix m n R} : + block_diagonal (∑' x, f x) = ∑' x, block_diagonal (f x) := +begin + by_cases hf : summable f, + { exact hf.has_sum.matrix_block_diagonal.tsum_eq.symm }, + { have hft := summable_matrix_block_diagonal.not.mpr hf, + rw [tsum_eq_zero_of_not_summable hf, tsum_eq_zero_of_not_summable hft], + exact block_diagonal_zero }, +end + +lemma has_sum.matrix_block_diag {f : X → matrix (m × p) (n × p) R} + {a : matrix (m × p) (n × p) R} (hf : has_sum f a) : + has_sum (λ x, block_diag (f x)) (block_diag a) := +(hf.map (block_diag_add_monoid_hom m n p R) $ continuous.matrix_block_diag continuous_id : _) + +lemma summable.matrix_block_diag {f : X → matrix (m × p) (n × p) R} (hf : summable f) : + summable (λ x, block_diag (f x)) := +hf.has_sum.matrix_block_diag.summable + lemma has_sum.matrix_block_diagonal' [decidable_eq l] {f : X → Π i, matrix (m' i) (n' i) R} {a : Π i, matrix (m' i) (n' i) R} (hf : has_sum f a) : has_sum (λ x, block_diagonal' (f x)) (block_diagonal' a) := (hf.map (block_diagonal'_add_monoid_hom m' n' R) $ continuous.matrix_block_diagonal' $ by exact continuous_id : _) -lemma summable.matrix_block_diagonal' [decidable_eq l] {f : X → Π i, matrix (m' i) (n' i) R} - (hf : summable f) : +lemma summable.matrix_block_diagonal' [decidable_eq l] + {f : X → Π i, matrix (m' i) (n' i) R} (hf : summable f) : summable (λ x, block_diagonal' (f x)) := hf.has_sum.matrix_block_diagonal'.summable +lemma summable_matrix_block_diagonal' [decidable_eq l] {f : X → Π i, matrix (m' i) (n' i) R} : + summable (λ x, block_diagonal' (f x)) ↔ summable f := +(summable.map_iff_of_left_inverse + (matrix.block_diagonal'_add_monoid_hom m' n' R) (matrix.block_diag'_add_monoid_hom m' n' R) + (by exact continuous.matrix_block_diagonal' continuous_id) + (by exact continuous.matrix_block_diag' continuous_id) + (λ A, block_diag'_block_diagonal' A) : _) + +lemma matrix.block_diagonal'_tsum [decidable_eq l] [t2_space R] + {f : X → Π i, matrix (m' i) (n' i) R} : + block_diagonal' (∑' x, f x) = ∑' x, block_diagonal' (f x) := +begin + by_cases hf : summable f, + { exact hf.has_sum.matrix_block_diagonal'.tsum_eq.symm }, + { have hft := summable_matrix_block_diagonal'.not.mpr hf, + rw [tsum_eq_zero_of_not_summable hf, tsum_eq_zero_of_not_summable hft], + exact block_diagonal'_zero }, +end + +lemma has_sum.matrix_block_diag' {f : X → matrix (Σ i, m' i) (Σ i, n' i) R} + {a : matrix (Σ i, m' i) (Σ i, n' i) R} (hf : has_sum f a) : + has_sum (λ x, block_diag' (f x)) (block_diag' a) := +(hf.map (block_diag'_add_monoid_hom m' n' R) $ continuous.matrix_block_diag' continuous_id : _) + +lemma summable.matrix_block_diag' {f : X → matrix (Σ i, m' i) (Σ i, n' i) R} (hf : summable f) : + summable (λ x, block_diag' (f x)) := +hf.has_sum.matrix_block_diag'.summable + end block_matrices end tsum From 0c9b72632bbe90c3e09adb1977b7e7210e213002 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Thu, 5 May 2022 03:36:27 +0000 Subject: [PATCH 130/153] feat(algebra/group/{pi, opposite}): add missing pi and opposite defs for `mul_hom` (#13956) The declaration names and the contents of these definitions are all copied from the corresponding ones for `monoid_hom`. --- src/algebra/group/opposite.lean | 58 +++++++++++++++++++++++++++++++++ src/algebra/group/pi.lean | 40 +++++++++++++++++++++++ 2 files changed, 98 insertions(+) diff --git a/src/algebra/group/opposite.lean b/src/algebra/group/opposite.lean index 1fe1527c194aa..f481a76f59aa5 100644 --- a/src/algebra/group/opposite.lean +++ b/src/algebra/group/opposite.lean @@ -228,6 +228,26 @@ def mul_equiv.inv' (G : Type*) [group G] : G ≃* Gᵐᵒᵖ := { map_mul' := λ x y, unop_injective $ mul_inv_rev x y, .. (equiv.inv G).trans op_equiv } +/-- A semigroup homomorphism `f : M →ₙ* N` such that `f x` commutes with `f y` for all `x, y` +defines a semigroup homomorphism to `Nᵐᵒᵖ`. -/ +@[to_additive "An additive semigroup homomorphism `f : add_hom M N` such that `f x` additively +commutes with `f y` for all `x, y` defines an additive semigroup homomorphism to `Sᵃᵒᵖ`.", + simps {fully_applied := ff}] +def mul_hom.to_opposite {M N : Type*} [has_mul M] [has_mul N] (f : M →ₙ* N) + (hf : ∀ x y, commute (f x) (f y)) : M →ₙ* Nᵐᵒᵖ := +{ to_fun := mul_opposite.op ∘ f, + map_mul' := λ x y, by simp [(hf x y).eq] } + +/-- A semigroup homomorphism `f : M →ₙ* N` such that `f x` commutes with `f y` for all `x, y` +defines a semigroup homomorphism from `Mᵐᵒᵖ`. -/ +@[to_additive "An additive semigroup homomorphism `f : add_hom M N` such that `f x` additively +commutes with `f y` for all `x`, `y` defines an additive semigroup homomorphism from `Mᵃᵒᵖ`.", + simps {fully_applied := ff}] +def mul_hom.from_opposite {M N : Type*} [has_mul M] [has_mul N] (f : M →ₙ* N) + (hf : ∀ x y, commute (f x) (f y)) : Mᵐᵒᵖ →ₙ* N := +{ to_fun := f ∘ mul_opposite.unop, + map_mul' := λ x y, (f.map_mul _ _).trans (hf _ _).eq } + /-- A monoid homomorphism `f : M →* N` such that `f x` commutes with `f y` for all `x, y` defines a monoid homomorphism to `Nᵐᵒᵖ`. -/ @[to_additive "An additive monoid homomorphism `f : M →+ N` such that `f x` additively commutes @@ -270,6 +290,44 @@ lemma units.coe_op_equiv_symm {M} [monoid M] (u : (Mˣ)ᵐᵒᵖ) : (units.op_equiv.symm u : Mᵐᵒᵖ) = op (u.unop : M) := rfl +/-- A semigroup homomorphism `M →ₙ* N` can equivalently be viewed as a semigroup homomorphism +`Mᵐᵒᵖ →ₙ* Nᵐᵒᵖ`. This is the action of the (fully faithful) `ᵐᵒᵖ`-functor on morphisms. -/ +@[to_additive "An additive semigroup homomorphism `add_hom M N` can equivalently be viewed as an +additive semigroup homomorphism `add_hom Mᵃᵒᵖ Nᵃᵒᵖ`. This is the action of the (fully faithful) +`ᵃᵒᵖ`-functor on morphisms.", simps] +def mul_hom.op {M N} [has_mul M] [has_mul N] : + (M →ₙ* N) ≃ (Mᵐᵒᵖ →ₙ* Nᵐᵒᵖ) := +{ to_fun := λ f, { to_fun := op ∘ f ∘ unop, + map_mul' := λ x y, unop_injective (f.map_mul y.unop x.unop) }, + inv_fun := λ f, { to_fun := unop ∘ f ∘ op, + map_mul' := λ x y, congr_arg unop (f.map_mul (op y) (op x)) }, + left_inv := λ f, by { ext, refl }, + right_inv := λ f, by { ext x, simp } } + +/-- The 'unopposite' of a semigroup homomorphism `Mᵐᵒᵖ →ₙ* Nᵐᵒᵖ`. Inverse to `mul_hom.op`. -/ +@[simp, to_additive "The 'unopposite' of an additive semigroup homomorphism `Mᵃᵒᵖ →ₙ+ Nᵃᵒᵖ`. Inverse +to `add_hom.op`."] +def mul_hom.unop {M N} [has_mul M] [has_mul N] : + (Mᵐᵒᵖ →ₙ* Nᵐᵒᵖ) ≃ (M →ₙ* N) := mul_hom.op.symm + +/-- An additive semigroup homomorphism `add_hom M N` can equivalently be viewed as an additive +homomorphism `add_hom Mᵐᵒᵖ Nᵐᵒᵖ`. This is the action of the (fully faithful) `ᵐᵒᵖ`-functor on +morphisms. -/ +@[simps] +def add_hom.mul_op {M N} [has_add M] [has_add N] : + (add_hom M N) ≃ (add_hom Mᵐᵒᵖ Nᵐᵒᵖ) := +{ to_fun := λ f, { to_fun := op ∘ f ∘ unop, + map_add' := λ x y, unop_injective (f.map_add x.unop y.unop) }, + inv_fun := λ f, { to_fun := unop ∘ f ∘ op, + map_add' := λ x y, congr_arg unop (f.map_add (op x) (op y)) }, + left_inv := λ f, by { ext, refl }, + right_inv := λ f, by { ext, simp } } + +/-- The 'unopposite' of an additive semigroup hom `αᵐᵒᵖ →+ βᵐᵒᵖ`. Inverse to +`add_hom.mul_op`. -/ +@[simp] def add_hom.mul_unop {α β} [has_add α] [has_add β] : + (add_hom αᵐᵒᵖ βᵐᵒᵖ) ≃ (add_hom α β) := add_hom.mul_op.symm + /-- A monoid homomorphism `M →* N` can equivalently be viewed as a monoid homomorphism `Mᵐᵒᵖ →* Nᵐᵒᵖ`. This is the action of the (fully faithful) `ᵐᵒᵖ`-functor on morphisms. -/ @[to_additive "An additive monoid homomorphism `M →+ N` can equivalently be viewed as an diff --git a/src/algebra/group/pi.lean b/src/algebra/group/pi.lean index 6b7c5f6285253..2732c4165b978 100644 --- a/src/algebra/group/pi.lean +++ b/src/algebra/group/pi.lean @@ -131,6 +131,46 @@ namespace mul_hom end mul_hom +section mul_hom + +variables (f) [Π i, has_mul (f i)] + +/-- Evaluation of functions into an indexed collection of semigroups at a point is a semigroup +homomorphism. +This is `function.eval i` as a `mul_hom`. -/ +@[to_additive "Evaluation of functions into an indexed collection of additive semigroups at a +point is an additive semigroup homomorphism. +This is `function.eval i` as an `add_hom`.", simps] +def pi.eval_mul_hom (i : I) : (Π i, f i) →ₙ* f i := +{ to_fun := λ g, g i, + map_mul' := λ x y, pi.mul_apply _ _ i, } + +/-- `function.const` as a `mul_hom`. -/ +@[to_additive "`function.const` as an `add_hom`.", simps] +def pi.const_mul_hom (α β : Type*) [has_mul β] : β →ₙ* (α → β) := +{ to_fun := function.const α, + map_mul' := λ _ _, rfl } + +/-- Coercion of a `mul_hom` into a function is itself a `mul_hom`. +See also `mul_hom.eval`. -/ +@[to_additive "Coercion of an `add_hom` into a function is itself a `add_hom`. +See also `add_hom.eval`. ", simps] +def mul_hom.coe_fn (α β : Type*) [has_mul α] [comm_semigroup β] : (α →ₙ* β) →ₙ* (α → β) := +{ to_fun := λ g, g, + map_mul' := λ x y, rfl, } + +/-- Semigroup homomorphism between the function spaces `I → α` and `I → β`, induced by a semigroup +homomorphism `f` between `α` and `β`. -/ +@[to_additive "Additive semigroup homomorphism between the function spaces `I → α` and `I → β`, +induced by an additive semigroup homomorphism `f` between `α` and `β`", simps] +protected def mul_hom.comp_left {α β : Type*} [has_mul α] [has_mul β] (f : α →ₙ* β) + (I : Type*) : + (I → α) →ₙ* (I → β) := +{ to_fun := λ h, f ∘ h, + map_mul' := λ _ _, by ext; simp } + +end mul_hom + section monoid_hom variables (f) [Π i, mul_one_class (f i)] From b1da074a9113f2dde55f8c26329e58f0612f5e1b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 5 May 2022 05:36:19 +0000 Subject: [PATCH 131/153] feat(data/sum/basic): Shortcuts for the ternary sum of types (#13678) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define `sum3.in₀`, `sum3.in₁`, `sum3.in₂`, shortcut patterns for pattern-matching on a ternary sum of types. Co-authored-by: Bhavik Mehta --- src/data/sum/basic.lean | 21 +++++++++++++++++++-- 1 file changed, 19 insertions(+), 2 deletions(-) diff --git a/src/data/sum/basic.lean b/src/data/sum/basic.lean index 17c12cd23bb47..5bcda17b5f4ed 100644 --- a/src/data/sum/basic.lean +++ b/src/data/sum/basic.lean @@ -325,10 +325,10 @@ have aca : ∀ a, acc (lex r s) (inl a), from λ a, lex_acc_inl (ha.apply a), end lex end sum -namespace function - open sum +namespace function + lemma injective.sum_elim {f : α → γ} {g : β → γ} (hf : injective f) (hg : injective g) (hfg : ∀ a b, f a ≠ g b) : injective (sum.elim f g) @@ -348,3 +348,20 @@ lemma surjective.sum_map {f : α → β} {g : α' → β'} (hf : surjective f) ( | (inr y) := let ⟨x, hx⟩ := hg y in ⟨inr x, congr_arg inr hx⟩ end function + +/-! +### Ternary sum + +Abbreviations for the maps from the summands to `α ⊕ β ⊕ γ`. This is useful for pattern-matching. +-/ + +namespace sum3 + +/-- The map from the first summand into a ternary sum. -/ +@[pattern, simp, reducible] def in₀ (a) : α ⊕ β ⊕ γ := inl a +/-- The map from the second summand into a ternary sum. -/ +@[pattern, simp, reducible] def in₁ (b) : α ⊕ β ⊕ γ := inr $ inl b +/-- The map from the third summand into a ternary sum. -/ +@[pattern, simp, reducible] def in₂ (c) : α ⊕ β ⊕ γ := inr $ inr c + +end sum3 From 524793de15bc4c52ee32d254e7d7867c7176b3af Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Thu, 5 May 2022 05:36:20 +0000 Subject: [PATCH 132/153] feat(representation_theory): Action V G is rigid whenever V is (#13738) Co-authored-by: Scott Morrison Co-authored-by: Johan Commelin --- src/algebra/category/Group/basic.lean | 9 +++++++++ src/algebra/category/Mon/basic.lean | 3 +++ src/representation_theory/Action.lean | 28 +++++++++++++++++++++++++-- 3 files changed, 38 insertions(+), 2 deletions(-) diff --git a/src/algebra/category/Group/basic.lean b/src/algebra/category/Group/basic.lean index 6791adc480959..17a1980b15b3d 100644 --- a/src/algebra/category/Group/basic.lean +++ b/src/algebra/category/Group/basic.lean @@ -74,6 +74,9 @@ by { ext1, apply w } @[to_additive has_forget_to_AddMon] instance has_forget_to_Mon : has_forget₂ Group Mon := bundled_hom.forget₂ _ _ +@[to_additive] instance : has_coe Group.{u} Mon.{u} := +{ coe := (forget₂ Group Mon).obj, } + end Group /-- The category of commutative groups and group morphisms. -/ @@ -134,10 +137,16 @@ by { ext1, apply w } @[to_additive has_forget_to_AddGroup] instance has_forget_to_Group : has_forget₂ CommGroup Group := bundled_hom.forget₂ _ _ +@[to_additive] instance : has_coe CommGroup.{u} Group.{u} := +{ coe := (forget₂ CommGroup Group).obj, } + @[to_additive has_forget_to_AddCommMon] instance has_forget_to_CommMon : has_forget₂ CommGroup CommMon := induced_category.has_forget₂ (λ G : CommGroup, CommMon.of G) +@[to_additive] instance : has_coe CommGroup.{u} CommMon.{u} := +{ coe := (forget₂ CommGroup CommMon).obj, } + end CommGroup -- This example verifies an improvement possible in Lean 3.8. diff --git a/src/algebra/category/Mon/basic.lean b/src/algebra/category/Mon/basic.lean index 0cb4c2482a989..198b77989a8e7 100644 --- a/src/algebra/category/Mon/basic.lean +++ b/src/algebra/category/Mon/basic.lean @@ -118,6 +118,9 @@ instance (M : CommMon) : comm_monoid M := M.str @[to_additive has_forget_to_AddMon] instance has_forget_to_Mon : has_forget₂ CommMon Mon := bundled_hom.forget₂ _ _ +@[to_additive] instance : has_coe CommMon.{u} Mon.{u} := +{ coe := (forget₂ CommMon Mon).obj, } + end CommMon -- We verify that the coercions of morphisms to functions work correctly: diff --git a/src/representation_theory/Action.lean b/src/representation_theory/Action.lean index 270fd0ad25b90..bb29fd03d19cc 100644 --- a/src/representation_theory/Action.lean +++ b/src/representation_theory/Action.lean @@ -10,6 +10,8 @@ import category_theory.limits.preserves.basic import category_theory.adjunction.limits import category_theory.monoidal.functor_category import category_theory.monoidal.transport +import category_theory.monoidal.rigid.of_equivalence +import category_theory.monoidal.rigid.functor_category import category_theory.monoidal.linear import category_theory.monoidal.braided import category_theory.abelian.functor_category @@ -394,16 +396,38 @@ end instance [symmetric_category V] : symmetric_category (Action V G) := symmetric_category_of_faithful (forget_braided V G) -variables [preadditive V] [monoidal_preadditive V] - +section local attribute [simp] monoidal_preadditive.tensor_add monoidal_preadditive.add_tensor +variables [preadditive V] [monoidal_preadditive V] + instance : monoidal_preadditive (Action V G) := {} variables {R : Type*} [semiring R] [linear R V] [monoidal_linear R V] instance : monoidal_linear R (Action V G) := {} +end + +variables (V G) +noncomputable theory + +/-- Upgrading the functor `Action V G ⥤ (single_obj G ⥤ V)` to a monoidal functor. -/ +def functor_category_monoidal_equivalence : monoidal_functor (Action V G) (single_obj G ⥤ V) := +monoidal.from_transported (Action.functor_category_equivalence _ _).symm + +instance : is_equivalence ((functor_category_monoidal_equivalence V G).to_functor) := +by { change is_equivalence (Action.functor_category_equivalence _ _).functor, apply_instance, } + +variables (H : Group.{u}) + +instance [rigid_category V] : rigid_category (single_obj (H : Mon.{u}) ⥤ V) := +by { change rigid_category (single_obj H ⥤ V), apply_instance } + +/-- If `V` is rigid, so is `Action V G`. -/ +instance [rigid_category V] : rigid_category (Action V H) := +rigid_category_of_equivalence (functor_category_monoidal_equivalence V _) + end monoidal /-- Actions/representations of the trivial group are just objects in the ambient category. -/ From 116435d6081386854378a7f31333f05f9932cf24 Mon Sep 17 00:00:00 2001 From: Alex J Best Date: Thu, 5 May 2022 05:36:21 +0000 Subject: [PATCH 133/153] feat(tactic/alias): fix alias docstrings for implications from iffs (#13944) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Now they say for instance: ```lean le_inv_mul_of_mul_le : ∀ {α : Type u} [_inst_1 : group α] [_inst_2 : has_le α] [_inst_3 : covariant_class α α has_mul.mul has_le.le] {a b c : α}, a * b ≤ c → b ≤ a⁻¹ * c **Alias** of the reverse direction of `le_inv_mul_iff_mul_le`. ``` https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/.60alias.60.20issue.20in.20algebra.2Eorder.2Egroup/near/281158569 --- src/tactic/alias.lean | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/src/tactic/alias.lean b/src/tactic/alias.lean index a48ebe2afc3b6..0ade90b333a43 100644 --- a/src/tactic/alias.lean +++ b/src/tactic/alias.lean @@ -129,20 +129,20 @@ input theorem has the form `A_iff_B` or `A_iff_B_left` etc. do old ← ident, d ← (do old ← resolve_constant old, get_decl old) <|> fail ("declaration " ++ to_string old ++ " not found"), - let doc := λ al : name, meta_info.doc_string.get_or_else $ - "**Alias** of `" ++ to_string old ++ "`.", + let doc := λ (al : name) (inf : string), meta_info.doc_string.get_or_else $ + sformat!"**Alias** of {inf}`{old}`.", do { tk "←" <|> tk "<-", aliases ← many ident, - ↑(aliases.mmap' $ λ al, alias_direct d (doc al) al) } <|> + ↑(aliases.mmap' $ λ al, alias_direct d (doc al "") 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) left `iff.mp, - alias_iff d (doc right) right `iff.mpr } + alias_iff d (doc left "the forward direction of ") left `iff.mp, + alias_iff d (doc right "the reverse direction of ") right `iff.mpr } add_tactic_doc { name := "alias", From 63875eaefe9ba0828f5d050d4afb2744487ed0bc Mon Sep 17 00:00:00 2001 From: leanprover-community-bot Date: Thu, 5 May 2022 05:36:22 +0000 Subject: [PATCH 134/153] chore(scripts): update nolints.txt (#13964) I am happy to remove some nolints for you! --- scripts/nolints.txt | 16 ---------------- 1 file changed, 16 deletions(-) diff --git a/scripts/nolints.txt b/scripts/nolints.txt index ccb5f9922f946..3551a7d157a1f 100644 --- a/scripts/nolints.txt +++ b/scripts/nolints.txt @@ -27,26 +27,10 @@ apply_nolint mul_hom.ext_iff to_additive_doc apply_nolint one_hom.comp_assoc to_additive_doc -- algebra/order/group.lean -apply_nolint inv_le_of_inv_le' to_additive_doc -apply_nolint inv_lt_of_inv_lt' to_additive_doc -apply_nolint inv_mul_lt_of_lt_mul to_additive_doc -apply_nolint inv_of_one_lt_inv to_additive_doc -apply_nolint le_inv_mul_of_mul_le to_additive_doc -apply_nolint le_inv_of_le_inv to_additive_doc -apply_nolint le_one_of_one_le_inv to_additive_doc apply_nolint left.inv_le_one_iff to_additive_doc apply_nolint left.inv_lt_one_iff to_additive_doc apply_nolint left.one_le_inv_iff to_additive_doc apply_nolint left.one_lt_inv_iff to_additive_doc -apply_nolint lt_inv_mul_of_mul_lt to_additive_doc -apply_nolint lt_inv_of_lt_inv to_additive_doc -apply_nolint lt_mul_of_inv_mul_lt to_additive_doc -apply_nolint lt_of_inv_lt_inv to_additive_doc -apply_nolint mul_le_of_le_inv_mul to_additive_doc -apply_nolint mul_lt_of_lt_inv_mul to_additive_doc -apply_nolint one_le_of_inv_le_one to_additive_doc -apply_nolint one_lt_inv_of_inv to_additive_doc -apply_nolint one_lt_of_inv_lt_one to_additive_doc apply_nolint right.inv_le_one_iff to_additive_doc apply_nolint right.one_le_inv_iff to_additive_doc From 03fbe7de7be7416228b491873db7946654370b52 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Martin=20Dvo=C5=99=C3=A1k?= Date: Thu, 5 May 2022 05:36:23 +0000 Subject: [PATCH 135/153] Update CODE_OF_CONDUCT.md (#13965) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit deleted one character (duplicate space) Co-authored-by: Martin Dvořák --- CODE_OF_CONDUCT.md | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/CODE_OF_CONDUCT.md b/CODE_OF_CONDUCT.md index a12ebc6fe6930..16a5c29a35aeb 100644 --- a/CODE_OF_CONDUCT.md +++ b/CODE_OF_CONDUCT.md @@ -108,7 +108,7 @@ Violating these terms may lead to a permanent ban. ### 4. Permanent Ban **Community Impact**: Demonstrating a pattern of violation of community -standards, including sustained inappropriate behavior, harassment of an +standards, including sustained inappropriate behavior, harassment of an individual, or aggression toward or disparagement of classes of individuals. **Consequence**: A permanent ban from any sort of public interaction within From 7eacca31fa6f5980038bf433fc8e75a4df6ac803 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 5 May 2022 07:41:33 +0000 Subject: [PATCH 136/153] =?UTF-8?q?feat(analysis/normed/normed=5Ffield):?= =?UTF-8?q?=20limit=20of=20`=E2=88=A5a=20*=20x=E2=88=A5`=20as=20`=E2=88=A5?= =?UTF-8?q?x=E2=88=A5=20=E2=86=92=20=E2=88=9E`=20(#13819)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit These lemmas should use `bornology.cobounded` but we don't have an instance `pseudo_metric_space α -> bornology α` yet. --- src/analysis/normed/group/basic.lean | 6 ++++++ src/analysis/normed/normed_field.lean | 16 ++++++++++++++-- 2 files changed, 20 insertions(+), 2 deletions(-) diff --git a/src/analysis/normed/group/basic.lean b/src/analysis/normed/group/basic.lean index fe3da3b684567..05c587581b6df 100644 --- a/src/analysis/normed/group/basic.lean +++ b/src/analysis/normed/group/basic.lean @@ -166,6 +166,12 @@ by rw [sub_eq_add_neg, dist_self_add_right, norm_neg] @[simp] theorem dist_self_sub_left (g h : E) : dist (g - h) g = ∥h∥ := by rw [dist_comm, dist_self_sub_right] +/-- In a (semi)normed group, negation `x ↦ -x` tends to infinity at infinity. TODO: use +`bornology.cobounded` instead of `filter.comap has_norm.norm filter.at_top`. -/ +lemma filter.tendsto_neg_cobounded : + tendsto (has_neg.neg : E → E) (comap norm at_top) (comap norm at_top) := +by simpa only [norm_neg, tendsto_comap_iff, (∘)] using tendsto_comap + /-- **Triangle inequality** for the norm. -/ lemma norm_add_le (g h : E) : ∥g + h∥ ≤ ∥g∥ + ∥h∥ := by simpa [dist_eq_norm] using dist_triangle g 0 (-h) diff --git a/src/analysis/normed/normed_field.lean b/src/analysis/normed/normed_field.lean index 2247519ba7780..ba363db236272 100644 --- a/src/analysis/normed/normed_field.lean +++ b/src/analysis/normed/normed_field.lean @@ -385,8 +385,6 @@ instance normed_division_ring.to_norm_one_class : norm_one_class α := ⟨mul_left_cancel₀ (mt norm_eq_zero.1 (@one_ne_zero α _ _)) $ by rw [← norm_mul, mul_one, mul_one]⟩ -export norm_one_class (norm_one) - @[simp] lemma nnnorm_mul (a b : α) : ∥a * b∥₊ = ∥a∥₊ * ∥b∥₊ := nnreal.eq $ norm_mul a b @@ -422,6 +420,20 @@ nnreal.eq $ by simp @[simp] lemma nnnorm_zpow : ∀ (a : α) (n : ℤ), ∥a ^ n∥₊ = ∥a∥₊ ^ n := (nnnorm_hom : α →*₀ ℝ≥0).map_zpow +/-- Multiplication on the left by a nonzero element of a normed division ring tends to infinity at +infinity. TODO: use `bornology.cobounded` instead of `filter.comap has_norm.norm filter.at_top`. -/ +lemma filter.tendsto_mul_left_cobounded {a : α} (ha : a ≠ 0) : + tendsto ((*) a) (comap norm at_top) (comap norm at_top) := +by simpa only [tendsto_comap_iff, (∘), norm_mul] + using tendsto_const_nhds.mul_at_top (norm_pos_iff.2 ha) tendsto_comap + +/-- Multiplication on the right by a nonzero element of a normed division ring tends to infinity at +infinity. TODO: use `bornology.cobounded` instead of `filter.comap has_norm.norm filter.at_top`. -/ +lemma filter.tendsto_mul_right_cobounded {a : α} (ha : a ≠ 0) : + tendsto (λ x, x * a) (comap norm at_top) (comap norm at_top) := +by simpa only [tendsto_comap_iff, (∘), norm_mul] + using tendsto_comap.at_top_mul (norm_pos_iff.2 ha) tendsto_const_nhds + @[priority 100] -- see Note [lower instance priority] instance normed_division_ring.to_has_continuous_inv₀ : has_continuous_inv₀ α := begin From 4f7603c797564cc02a80feec39cd591c9921f1fa Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 5 May 2022 07:41:35 +0000 Subject: [PATCH 137/153] chore(data/matrix/basic): add more lemmas about `conj_transpose` and `smul` (#13938) Unfortunately the `star_module` typeclass is of no help here; see [this Zulip thread](https://leanprover.zulipchat.com/#narrow/stream/116395-maths/topic/Is.20star_module.20sensible.20for.20non-commutative.20rings.3F/near/257272767) for some discussion. In the meantime, this adds the lemmas for the most frequent special cases. --- src/algebra/star/basic.lean | 4 ++-- src/algebra/star/module.lean | 4 ++++ src/data/matrix/basic.lean | 35 +++++++++++++++++++++++++++++++++-- 3 files changed, 39 insertions(+), 4 deletions(-) diff --git a/src/algebra/star/basic.lean b/src/algebra/star/basic.lean index f4e19a3065b54..0b2482caa8732 100644 --- a/src/algebra/star/basic.lean +++ b/src/algebra/star/basic.lean @@ -217,11 +217,11 @@ star_eq_zero.not star (r - s) = star r - star s := (star_add_equiv : R ≃+ R).map_sub _ _ -@[simp] lemma star_nsmul [add_comm_monoid R] [star_add_monoid R] (x : R) (n : ℕ) : +@[simp] lemma star_nsmul [add_monoid R] [star_add_monoid R] (x : R) (n : ℕ) : star (n • x) = n • star x := (star_add_equiv : R ≃+ R).to_add_monoid_hom.map_nsmul _ _ -@[simp] lemma star_zsmul [add_comm_group R] [star_add_monoid R] (x : R) (n : ℤ) : +@[simp] lemma star_zsmul [add_group R] [star_add_monoid R] (x : R) (n : ℤ) : star (n • x) = n • star x := (star_add_equiv : R ≃+ R).to_add_monoid_hom.map_zsmul _ _ diff --git a/src/algebra/star/module.lean b/src/algebra/star/module.lean index eaa1454884a1f..d0b67bf4419b4 100644 --- a/src/algebra/star/module.lean +++ b/src/algebra/star/module.lean @@ -25,6 +25,10 @@ It is defined on a star algebra `A` over the base ring `R`. equivalence. -/ +@[simp] lemma star_rat_smul {R : Type*} [add_comm_group R] [star_add_monoid R] [module ℚ R] + (x : R) (n : ℚ) : star (n • x) = n • star x := +map_rat_smul (star_add_equiv : R ≃+ R) _ _ + /-- If `A` is a module over a commutative `R` with compatible actions, then `star` is a semilinear equivalence. -/ @[simps] diff --git a/src/data/matrix/basic.lean b/src/data/matrix/basic.lean index a3dd7ba1f9589..e5d4a6f4d63bd 100644 --- a/src/data/matrix/basic.lean +++ b/src/data/matrix/basic.lean @@ -9,6 +9,7 @@ import algebra.big_operators.ring import algebra.module.linear_map import algebra.module.pi import algebra.ring.equiv +import algebra.star.module import algebra.star.pi import data.fintype.card @@ -1334,8 +1335,38 @@ matrix.ext $ by simp (M - N)ᴴ = Mᴴ - Nᴴ := matrix.ext $ by simp -@[simp] lemma conj_transpose_smul [comm_semigroup α] [star_semigroup α] (c : α) (M : matrix m n α) : - (c • M)ᴴ = (star c) • Mᴴ := +/-- Note that `star_module` is quite a strong requirement; as such we also provide the following +variants which this lemma would not apply to: +* `matrix.conj_transpose_smul_non_comm` +* `matrix.conj_transpose_nsmul` +* `matrix.conj_transpose_zsmul` +* `matrix.conj_transpose_rat_smul` +-/ +@[simp] lemma conj_transpose_smul [has_star R] [has_star α] [has_scalar R α] [star_module R α] + (c : R) (M : matrix m n α) : + (c • M)ᴴ = star c • Mᴴ := +matrix.ext $ λ i j, star_smul _ _ + +@[simp] lemma conj_transpose_smul_non_comm [has_star R] [has_star α] + [has_scalar R α] [has_scalar Rᵐᵒᵖ α] (c : R) (M : matrix m n α) + (h : ∀ (r : R) (a : α), star (r • a) = mul_opposite.op (star r) • star a) : + (c • M)ᴴ = mul_opposite.op (star c) • Mᴴ := +matrix.ext $ by simp [h] + +@[simp] lemma conj_transpose_smul_self [semigroup α] [star_semigroup α] (c : α) + (M : matrix m n α) : (c • M)ᴴ = mul_opposite.op (star c) • Mᴴ := +conj_transpose_smul_non_comm c M star_mul + +@[simp] lemma conj_transpose_nsmul [add_monoid α] [star_add_monoid α] (c : ℕ) (M : matrix m n α) : + (c • M)ᴴ = c • Mᴴ := +matrix.ext $ by simp + +@[simp] lemma conj_transpose_zsmul [add_group α] [star_add_monoid α] (c : ℤ) (M : matrix m n α) : + (c • M)ᴴ = c • Mᴴ := +matrix.ext $ by simp + +@[simp] lemma conj_transpose_rat_smul [add_comm_group α] [star_add_monoid α] [module ℚ α] (c : ℚ) + (M : matrix m n α) : (c • M)ᴴ = c • Mᴴ := matrix.ext $ by simp @[simp] lemma conj_transpose_mul [fintype n] [non_unital_semiring α] [star_ring α] From bd944fe1f6f26ea22243d429b905ae3fd0b1b94f Mon Sep 17 00:00:00 2001 From: Alex J Best Date: Thu, 5 May 2022 07:41:36 +0000 Subject: [PATCH 138/153] chore(linear_algebra/free_module): fix name in doc (#13942) --- src/linear_algebra/free_module/basic.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/linear_algebra/free_module/basic.lean b/src/linear_algebra/free_module/basic.lean index 243ccdabbde6a..3550663fab514 100644 --- a/src/linear_algebra/free_module/basic.lean +++ b/src/linear_algebra/free_module/basic.lean @@ -65,11 +65,11 @@ section semiring variables (R M) [semiring R] [add_comm_monoid M] [module R M] [module.free R M] variables [add_comm_monoid N] [module R N] -/-- If `[finite_free R M]` then `choose_basis_index R M` is the `ι` which indexes the basis +/-- If `module.free R M` then `choose_basis_index R M` is the `ι` which indexes the basis `ι → M`. -/ @[nolint has_inhabited_instance] def choose_basis_index := (exists_basis R M).some.1 -/-- If `[finite_free R M]` then `choose_basis : ι → M` is the basis. +/-- If `module.free R M` then `choose_basis : ι → M` is the basis. Here `ι = choose_basis_index R M`. -/ noncomputable def choose_basis : basis (choose_basis_index R M) R M := (exists_basis R M).some.2 From 697012956aff5aead72f5287d61b62ff1123f2a2 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 5 May 2022 07:41:37 +0000 Subject: [PATCH 139/153] chore(algebra/group/units): add a lemma about is_unit on a coerced unit (#13947) --- src/algebra/group/units.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/src/algebra/group/units.lean b/src/algebra/group/units.lean index b20ec07a4a973..f9b8e3afd2df2 100644 --- a/src/algebra/group/units.lean +++ b/src/algebra/group/units.lean @@ -379,6 +379,10 @@ an additive monoid which is an additive unit."] noncomputable def is_unit.unit [monoid M] {a : M} (h : is_unit a) : Mˣ := (classical.some h).copy a (classical.some_spec h).symm _ rfl +@[simp, to_additive] +lemma is_unit.unit_of_coe_units [monoid M] {a : Mˣ} (h : is_unit (a : M)) : h.unit = a := +units.ext $ rfl + @[to_additive] lemma is_unit.unit_spec [monoid M] {a : M} (h : is_unit a) : ↑h.unit = a := rfl From f82067179d1ff74303c01bfdac6d13a967741aac Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 5 May 2022 07:41:38 +0000 Subject: [PATCH 140/153] ci(gitpod): do not rerun get-cache if a workspace is reloaded (#13949) Instead, only run it at workspace start. This prevents it clobbering local builds created with `lean --make src` or similar. I have no idea why the `. /home/gitpod/.profile` line is there, so I've left it to run in the same phase as before. --- .gitpod.yml | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/.gitpod.yml b/.gitpod.yml index 5208d63e74092..c0cac26eda069 100644 --- a/.gitpod.yml +++ b/.gitpod.yml @@ -6,4 +6,5 @@ vscode: - jroesch.lean tasks: - - command: . /home/gitpod/.profile && leanproject get-cache --fallback=download-all + - init: leanpkg config && leanproject get-cache --fallback=download-all + command: . /home/gitpod/.profile From 4dfbcac38fb3f7000760c585227c214fa0113fc9 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Thu, 5 May 2022 09:52:28 +0000 Subject: [PATCH 141/153] feat({data/{finset,set},order/filter}/pointwise): More basic API (#13899) More basic lemmas about pointwise operations on `set`/`finset`/`filter`. Also make the three APIs more consistent with each other. --- src/data/finset/basic.lean | 5 + src/data/finset/pointwise.lean | 144 +++++++++++------ src/data/set/basic.lean | 9 ++ src/data/set/pointwise.lean | 212 +++++++++++-------------- src/order/filter/n_ary.lean | 3 + src/order/filter/pointwise.lean | 48 ++++-- src/topology/algebra/filter_basis.lean | 2 +- 7 files changed, 243 insertions(+), 180 deletions(-) diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index 093d89db59223..f918904435070 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -2057,6 +2057,11 @@ theorem image_union [decidable_eq α] {f : α → β} (s₁ s₂ : finset α) : ext $ λ _, by simp only [mem_image, mem_union, exists_prop, or_and_distrib_right, exists_or_distrib] +lemma image_inter_subset [decidable_eq α] (f : α → β) (s t : finset α) : + (s ∩ t).image f ⊆ s.image f ∩ t.image f := +subset_inter (image_subset_image $ inter_subset_left _ _) $ + image_subset_image $ inter_subset_right _ _ + lemma image_inter [decidable_eq α] (s₁ s₂ : finset α) (hf : ∀ x y, f x = f y → x = y) : (s₁ ∩ s₂).image f = s₁.image f ∩ s₂.image f := ext $ by simp only [mem_image, exists_prop, mem_inter]; exact λ b, diff --git a/src/data/finset/pointwise.lean b/src/data/finset/pointwise.lean index ac275fb7102c7..dd2f87783c0ab 100644 --- a/src/data/finset/pointwise.lean +++ b/src/data/finset/pointwise.lean @@ -155,23 +155,32 @@ lemma coe_mul (s t : finset α) : (↑(s * t) : set α) = ↑s * ↑t := coe_ima @[simp, to_additive] lemma empty_mul (s : finset α) : ∅ * s = ∅ := image₂_empty_left @[simp, to_additive] lemma mul_empty (s : finset α) : s * ∅ = ∅ := image₂_empty_right - -@[simp, to_additive] -lemma mul_nonempty_iff (s t : finset α) : (s * t).nonempty ↔ s.nonempty ∧ t.nonempty := +@[simp, to_additive] lemma mul_eq_empty : s * t = ∅ ↔ s = ∅ ∨ t = ∅ := image₂_eq_empty_iff +@[simp, to_additive] lemma mul_nonempty : (s * t).nonempty ↔ s.nonempty ∧ t.nonempty := image₂_nonempty_iff - @[to_additive] lemma nonempty.mul : s.nonempty → t.nonempty → (s * t).nonempty := nonempty.image₂ +@[to_additive] lemma nonempty.of_mul_left : (s * t).nonempty → s.nonempty := nonempty.of_image₂_left +@[to_additive] lemma nonempty.of_mul_right : (s * t).nonempty → t.nonempty := +nonempty.of_image₂_right +@[simp, to_additive] lemma mul_singleton (a : α) : s * {a} = s.image (* a) := image₂_singleton_right +@[simp, to_additive] lemma singleton_mul (a : α) : {a} * s = s.image ((*) a) := +image₂_singleton_left +@[simp, to_additive] lemma singleton_mul_singleton (a b : α) : ({a} : finset α) * {b} = {a * b} := +image₂_singleton @[to_additive, mono] lemma mul_subset_mul : s₁ ⊆ s₂ → t₁ ⊆ t₂ → s₁ * t₁ ⊆ s₂ * t₂ := image₂_subset +@[to_additive] lemma mul_subset_mul_left : t₁ ⊆ t₂ → s * t₁ ⊆ s * t₂ := image₂_subset_left +@[to_additive] lemma mul_subset_mul_right : s₁ ⊆ s₂ → s₁ * t ⊆ s₂ * t := image₂_subset_right +@[to_additive] lemma mul_subset_iff : s * t ⊆ u ↔ ∀ (x ∈ s) (y ∈ t), x * y ∈ u := image₂_subset_iff attribute [mono] add_subset_add -@[simp, to_additive] lemma mul_singleton (a : α) : s * {a} = s.image (* a) := image₂_singleton_right -@[simp, to_additive] lemma singleton_mul (a : α) : {a} * s = s.image ((*) a) := -image₂_singleton_left - -@[simp, to_additive] -lemma singleton_mul_singleton (a b : α) : ({a} : finset α) * {b} = {a * b} := image₂_singleton +@[to_additive] lemma union_mul : (s₁ ∪ s₂) * t = s₁ * t ∪ s₂ * t := image₂_union_left +@[to_additive] lemma mul_union : s * (t₁ ∪ t₂) = s * t₁ ∪ s * t₂ := image₂_union_right +@[to_additive] lemma inter_mul_subset : (s₁ ∩ s₂) * t ⊆ s₁ * t ∩ (s₂ * t) := +image₂_inter_subset_left +@[to_additive] lemma mul_inter_subset : s * (t₁ ∩ t₂) ⊆ s * t₁ ∩ (s * t₂) := +image₂_inter_subset_right /-- If a finset `u` is contained in the product of two sets `s * t`, we can find two finsets `s'`, `t'` such that `s' ⊆ s`, `t' ⊆ t` and `u ⊆ s' * t'`. -/ @@ -282,23 +291,32 @@ lemma coe_div (s t : finset α) : (↑(s / t) : set α) = ↑s / ↑t := coe_ima @[simp, to_additive] lemma empty_div (s : finset α) : ∅ / s = ∅ := image₂_empty_left @[simp, to_additive] lemma div_empty (s : finset α) : s / ∅ = ∅ := image₂_empty_right - -@[simp, to_additive] -lemma div_nonempty_iff (s t : finset α) : (s / t).nonempty ↔ s.nonempty ∧ t.nonempty := +@[simp, to_additive] lemma div_eq_empty : s / t = ∅ ↔ s = ∅ ∨ t = ∅ := image₂_eq_empty_iff +@[simp, to_additive] lemma div_nonempty : (s / t).nonempty ↔ s.nonempty ∧ t.nonempty := image₂_nonempty_iff - @[to_additive] lemma nonempty.div : s.nonempty → t.nonempty → (s / t).nonempty := nonempty.image₂ - -@[to_additive, mono] lemma div_subset_div : s₁ ⊆ s₂ → t₁ ⊆ t₂ → s₁ / t₁ ⊆ s₂ / t₂ := image₂_subset - -attribute [mono] add_subset_add - +@[to_additive] lemma nonempty.of_div_left : (s / t).nonempty → s.nonempty := nonempty.of_image₂_left +@[to_additive] lemma nonempty.of_div_right : (s / t).nonempty → t.nonempty := +nonempty.of_image₂_right @[simp, to_additive] lemma div_singleton (a : α) : s / {a} = s.image (/ a) := image₂_singleton_right @[simp, to_additive] lemma singleton_div (a : α) : {a} / s = s.image ((/) a) := image₂_singleton_left +@[simp, to_additive] lemma singleton_div_singleton (a b : α) : ({a} : finset α) / {b} = {a / b} := +image₂_singleton -@[simp, to_additive] -lemma singleton_div_singleton (a b : α) : ({a} : finset α) / {b} = {a / b} := image₂_singleton +@[to_additive, mono] lemma div_subset_div : s₁ ⊆ s₂ → t₁ ⊆ t₂ → s₁ / t₁ ⊆ s₂ / t₂ := image₂_subset +@[to_additive] lemma div_subset_div_left : t₁ ⊆ t₂ → s / t₁ ⊆ s / t₂ := image₂_subset_left +@[to_additive] lemma div_subset_div_right : s₁ ⊆ s₂ → s₁ / t ⊆ s₂ / t := image₂_subset_right +@[to_additive] lemma div_subset_iff : s / t ⊆ u ↔ ∀ (x ∈ s) (y ∈ t), x / y ∈ u := image₂_subset_iff + +attribute [mono] sub_subset_sub + +@[to_additive] lemma union_div : (s₁ ∪ s₂) / t = s₁ / t ∪ s₂ / t := image₂_union_left +@[to_additive] lemma div_union : s / (t₁ ∪ t₂) = s / t₁ ∪ s / t₂ := image₂_union_right +@[to_additive] lemma inter_div_subset : (s₁ ∩ s₂) / t ⊆ s₁ / t ∩ (s₂ / t) := +image₂_inter_subset_left +@[to_additive] lemma div_inter_subset : s / (t₁ ∩ t₂) ⊆ s / t₁ ∩ (s / t₂) := +image₂_inter_subset_right /-- If a finset `u` is contained in the product of two sets `s / t`, we can find two finsets `s'`, `t'` such that `s' ⊆ s`, `t' ⊆ t` and `u ⊆ s' / t'`. -/ @@ -448,25 +466,36 @@ coe_image₂ _ _ _ @[simp, to_additive] lemma empty_smul (t : finset β) : (∅ : finset α) • t = ∅ := image₂_empty_left @[simp, to_additive] lemma smul_empty (s : finset α) : s • (∅ : finset β) = ∅ := image₂_empty_right - -@[simp, to_additive] -lemma smul_nonempty_iff : (s • t).nonempty ↔ s.nonempty ∧ t.nonempty := image₂_nonempty_iff - +@[simp, to_additive] lemma smul_eq_empty : s • t = ∅ ↔ s = ∅ ∨ t = ∅ := image₂_eq_empty_iff +@[simp, to_additive] lemma smul_nonempty_iff : (s • t).nonempty ↔ s.nonempty ∧ t.nonempty := +image₂_nonempty_iff @[to_additive] lemma nonempty.smul : s.nonempty → t.nonempty → (s • t).nonempty := nonempty.image₂ +@[to_additive] lemma nonempty.of_smul_left : (s • t).nonempty → s.nonempty := +nonempty.of_image₂_left +@[to_additive] lemma nonempty.of_smul_right : (s • t).nonempty → t.nonempty := +nonempty.of_image₂_right +@[simp, to_additive] lemma smul_singleton (b : β) : s • ({b} : finset β) = s.image (• b) := +image₂_singleton_right +@[simp, to_additive] lemma singleton_smul (a : α) : ({a} : finset α) • t = t.image ((•) a) := +image₂_singleton_left +@[simp, to_additive] lemma singleton_smul_singleton (a : α) (b : β) : + ({a} : finset α) • ({b} : finset β) = {a • b} := +image₂_singleton @[to_additive, mono] lemma smul_subset_smul : s₁ ⊆ s₂ → t₁ ⊆ t₂ → s₁ • t₁ ⊆ s₂ • t₂ := image₂_subset +@[to_additive] lemma smul_subset_smul_left : t₁ ⊆ t₂ → s • t₁ ⊆ s • t₂ := image₂_subset_left +@[to_additive] lemma smul_subset_smul_right : s₁ ⊆ s₂ → s₁ • t ⊆ s₂ • t := image₂_subset_right +@[to_additive] lemma smul_subset_iff : s • t ⊆ u ↔ ∀ (a ∈ s) (b ∈ t), a • b ∈ u := image₂_subset_iff -attribute [mono] add_subset_add - -@[simp, to_additive] -lemma smul_singleton (b : β) : s • ({b} : finset β) = s.image (• b) := image₂_singleton_right - -@[simp, to_additive] -lemma singleton_smul (a : α) : ({a} : finset α) • t = t.image ((•) a) := image₂_singleton_left +attribute [mono] vadd_subset_vadd -@[simp, to_additive] -lemma singleton_smul_singleton (a : α) (b : β) : ({a} : finset α) • ({b} : finset β) = {a • b} := -image₂_singleton +@[to_additive] lemma union_smul [decidable_eq α] : (s₁ ∪ s₂) • t = s₁ • t ∪ s₂ • t := +image₂_union_left +@[to_additive] lemma smul_union : s • (t₁ ∪ t₂) = s • t₁ ∪ s • t₂ := image₂_union_right +@[to_additive] lemma inter_smul_subset [decidable_eq α] : (s₁ ∩ s₂) • t ⊆ s₁ • t ∩ s₂ • t := +image₂_inter_subset_left +@[to_additive] lemma smul_inter_subset : s • (t₁ ∩ t₂) ⊆ s • t₁ ∩ s • t₂ := +image₂_inter_subset_right /-- If a finset `u` is contained in the scalar product of two sets `s • t`, we can find two finsets `s'`, `t'` such that `s' ⊆ s`, `t' ⊆ t` and `u ⊆ s' • t'`. -/ @@ -481,7 +510,8 @@ end has_scalar /-! ### Finset addition/multiplication -/ section has_vsub -variables [decidable_eq α] [has_vsub α β] {s s₁ s₂ t t₁ t₂ u : finset β} {a : α} {b c : β} +variables [decidable_eq α] [has_vsub α β] {s s₁ s₂ t t₁ t₂ : finset β} {u : finset α} {a : α} + {b c : β} include α /-- The pointwise product of two finsets `s` and `t`: `s -ᵥ t = {x -ᵥ y | x ∈ s, y ∈ t}`. -/ @@ -502,26 +532,37 @@ lemma vsub_card_le : (s -ᵥ t : finset α).card ≤ s.card * t.card := card_ima @[simp] lemma empty_vsub (t : finset β) : (∅ : finset β) -ᵥ t = ∅ := image₂_empty_left @[simp] lemma vsub_empty (s : finset β) : s -ᵥ (∅ : finset β) = ∅ := image₂_empty_right - -@[simp] lemma vsub_nonempty_iff : (s -ᵥ t : finset α).nonempty ↔ s.nonempty ∧ t.nonempty := +@[simp] lemma vsub_eq_empty : s -ᵥ t = ∅ ↔ s = ∅ ∨ t = ∅ := image₂_eq_empty_iff +@[simp] lemma vsub_nonempty : (s -ᵥ t : finset α).nonempty ↔ s.nonempty ∧ t.nonempty := image₂_nonempty_iff - lemma nonempty.vsub : s.nonempty → t.nonempty → (s -ᵥ t : finset α).nonempty := nonempty.image₂ - -@[mono] lemma vsub_subset_vsub : s₁ ⊆ s₂ → t₁ ⊆ t₂ → s₁ -ᵥ t₁ ⊆ s₂ -ᵥ t₂ := image₂_subset - +lemma nonempty.of_vsub_left : (s -ᵥ t : finset α).nonempty → s.nonempty := nonempty.of_image₂_left +lemma nonempty.of_vsub_right : (s -ᵥ t : finset α).nonempty → t.nonempty := nonempty.of_image₂_right @[simp] lemma vsub_singleton (b : β) : s -ᵥ ({b} : finset β) = s.image (-ᵥ b) := image₂_singleton_right - @[simp] lemma singleton_vsub (a : β) : ({a} : finset β) -ᵥ t = t.image ((-ᵥ) a) := image₂_singleton_left +@[simp] lemma singleton_vsub_singleton (a b : β) : ({a} : finset β) -ᵥ {b} = {a -ᵥ b} := +image₂_singleton + +@[mono] lemma vsub_subset_vsub : s₁ ⊆ s₂ → t₁ ⊆ t₂ → s₁ -ᵥ t₁ ⊆ s₂ -ᵥ t₂ := image₂_subset +lemma vsub_subset_vsub_left : t₁ ⊆ t₂ → s -ᵥ t₁ ⊆ s -ᵥ t₂ := image₂_subset_left +lemma vsub_subset_vsub_right : s₁ ⊆ s₂ → s₁ -ᵥ t ⊆ s₂ -ᵥ t := image₂_subset_right +lemma vsub_subset_iff : s -ᵥ t ⊆ u ↔ ∀ (x ∈ s) (y ∈ t), x -ᵥ y ∈ u := image₂_subset_iff + +section +variables [decidable_eq β] + +lemma union_vsub : (s₁ ∪ s₂) -ᵥ t = (s₁ -ᵥ t) ∪ (s₂ -ᵥ t) := image₂_union_left +lemma vsub_union : s -ᵥ (t₁ ∪ t₂) = (s -ᵥ t₁) ∪ (s -ᵥ t₂) := image₂_union_right +lemma inter_vsub_subset : (s₁ ∩ s₂) -ᵥ t ⊆ (s₁ -ᵥ t) ∩ (s₂ -ᵥ t) := image₂_inter_subset_left +lemma vsub_inter_subset : s -ᵥ (t₁ ∩ t₂) ⊆ (s -ᵥ t₁) ∩ (s -ᵥ t₂) := image₂_inter_subset_right -@[simp] -lemma singleton_vsub_singleton (a b : β) : ({a} : finset β) -ᵥ {b} = {a -ᵥ b} := image₂_singleton +end /-- If a finset `u` is contained in the pointwise subtraction of two sets `s -ᵥ t`, we can find two finsets `s'`, `t'` such that `s' ⊆ s`, `t' ⊆ t` and `u ⊆ s' -ᵥ t'`. -/ -lemma subset_vsub {s t : set β} {u : finset α} : +lemma subset_vsub {s t : set β} : ↑u ⊆ s -ᵥ t → ∃ s' t' : finset β, ↑s' ⊆ s ∧ ↑t' ⊆ t ∧ u ⊆ s' -ᵥ t' := subset_image₂ @@ -555,10 +596,9 @@ lemma coe_smul_finset (s : finset β) : (↑(a • s) : set β) = a • s := coe @[to_additive] lemma smul_finset_card_le : (a • s).card ≤ s.card := card_image_le @[simp, to_additive] lemma smul_finset_empty (a : α) : a • (∅ : finset β) = ∅ := image_empty _ - -@[simp, to_additive] -lemma smul_finset_nonempty_iff : (a • s).nonempty ↔ s.nonempty := nonempty.image_iff _ - +@[simp, to_additive] lemma smul_finset_eq_empty : a • s = ∅ ↔ s = ∅ := image_eq_empty +@[simp, to_additive] lemma smul_finset_nonempty : (a • s).nonempty ↔ s.nonempty := +nonempty.image_iff _ @[to_additive] lemma nonempty.smul_finset (hs : s.nonempty) : (a • s).nonempty := hs.image _ @[to_additive, mono] @@ -569,6 +609,10 @@ attribute [mono] add_subset_add @[simp, to_additive] lemma smul_finset_singleton (b : β) : a • ({b} : finset β) = {a • b} := image_singleton _ _ +@[to_additive] lemma smul_finset_union : a • (s₁ ∪ s₂) = a • s₁ ∪ a • s₂ := image_union _ _ +@[to_additive] lemma smul_finset_inter_subset : a • (s₁ ∩ s₂) ⊆ a • s₁ ∩ (a • s₂) := +image_inter_subset _ _ _ + end has_scalar open_locale pointwise diff --git a/src/data/set/basic.lean b/src/data/set/basic.lean index f89814854c229..42e5bb1f8a96f 100644 --- a/src/data/set/basic.lean +++ b/src/data/set/basic.lean @@ -840,6 +840,9 @@ begin simp [eq_singleton_iff_nonempty_unique_mem, hs, ne_empty_iff_nonempty.2 hs], end +lemma nonempty.subset_singleton_iff (h : s.nonempty) : s ⊆ {a} ↔ s = {a} := +subset_singleton_iff_eq.trans $ or_iff_right h.ne_empty + lemma ssubset_singleton_iff {s : set α} {x : α} : s ⊂ {x} ↔ s = ∅ := begin rw [ssubset_iff_subset_ne, subset_singleton_iff_eq, or_and_distrib_right, and_not_self, or_false, @@ -2381,6 +2384,12 @@ lemma nonempty.image2 : s.nonempty → t.nonempty → (image2 f s t).nonempty := @[simp] lemma image2_nonempty_iff : (image2 f s t).nonempty ↔ s.nonempty ∧ t.nonempty := ⟨λ ⟨_, a, b, ha, hb, _⟩, ⟨⟨a, ha⟩, b, hb⟩, λ h, h.1.image2 h.2⟩ +lemma nonempty.of_image2_left (h : (image2 f s t).nonempty) : s.nonempty := +(image2_nonempty_iff.1 h).1 + +lemma nonempty.of_image2_right (h : (image2 f s t).nonempty) : t.nonempty := +(image2_nonempty_iff.1 h).2 + @[simp] lemma image2_eq_empty_iff : image2 f s t = ∅ ↔ s = ∅ ∨ t = ∅ := by simp_rw [←not_nonempty_iff_eq_empty, image2_nonempty_iff, not_and_distrib] diff --git a/src/data/set/pointwise.lean b/src/data/set/pointwise.lean index ee48dce8436dd..a93c4e8c54ceb 100644 --- a/src/data/set/pointwise.lean +++ b/src/data/set/pointwise.lean @@ -72,23 +72,15 @@ protected def has_one : has_one (set α) := ⟨{1}⟩ localized "attribute [instance] set.has_one set.has_zero" in pointwise -@[to_additive] -lemma singleton_one : ({1} : set α) = 1 := rfl - -@[simp, to_additive] -lemma mem_one : a ∈ (1 : set α) ↔ a = 1 := iff.rfl - -@[to_additive] -lemma one_mem_one : (1 : α) ∈ (1 : set α) := eq.refl _ - -@[simp, to_additive] -lemma one_subset : 1 ⊆ s ↔ (1 : α) ∈ s := singleton_subset_iff - -@[to_additive] -lemma one_nonempty : (1 : set α).nonempty := ⟨1, rfl⟩ - -@[simp, to_additive] -lemma image_one {f : α → β} : f '' 1 = {f 1} := image_singleton +@[to_additive] lemma singleton_one : ({1} : set α) = 1 := rfl +@[simp, to_additive] lemma mem_one : a ∈ (1 : set α) ↔ a = 1 := iff.rfl +@[to_additive] lemma one_mem_one : (1 : α) ∈ (1 : set α) := eq.refl _ +@[simp, to_additive] lemma one_subset : 1 ⊆ s ↔ (1 : α) ∈ s := singleton_subset_iff +@[to_additive] lemma one_nonempty : (1 : set α).nonempty := ⟨1, rfl⟩ +@[simp, to_additive] lemma image_one {f : α → β} : f '' 1 = {f 1} := image_singleton +@[to_additive] lemma subset_one_iff_eq : s ⊆ 1 ↔ s = ∅ ∨ s = 1 := subset_singleton_iff_eq +@[to_additive] lemma nonempty.subset_one_iff (h : s.nonempty) : s ⊆ 1 ↔ s = 1 := +h.subset_singleton_iff end one @@ -115,48 +107,45 @@ lemma image2_mul : image2 has_mul.mul s t = s * t := rfl @[to_additive] lemma mem_mul : a ∈ s * t ↔ ∃ x y, x ∈ s ∧ y ∈ t ∧ x * y = a := iff.rfl -@[to_additive] -lemma mul_mem_mul (ha : a ∈ s) (hb : b ∈ t) : a * b ∈ s * t := mem_image2_of_mem ha hb - -@[to_additive] -lemma mul_subset_mul (h₁ : s₁ ⊆ t₁) (h₂ : s₂ ⊆ t₂) : s₁ * s₂ ⊆ t₁ * t₂ := image2_subset h₁ h₂ +@[to_additive] lemma mul_mem_mul : a ∈ s → b ∈ t → a * b ∈ s * t := mem_image2_of_mem @[to_additive add_image_prod] lemma image_mul_prod : (λ x : α × α, x.fst * x.snd) '' s ×ˢ t = s * t := image_prod _ @[simp, to_additive] lemma empty_mul : ∅ * s = ∅ := image2_empty_left @[simp, to_additive] lemma mul_empty : s * ∅ = ∅ := image2_empty_right - +@[simp, to_additive] lemma mul_eq_empty : s * t = ∅ ↔ s = ∅ ∨ t = ∅ := image2_eq_empty_iff +@[simp, to_additive] lemma mul_nonempty : (s * t).nonempty ↔ s.nonempty ∧ t.nonempty := +image2_nonempty_iff +@[to_additive] lemma nonempty.mul : s.nonempty → t.nonempty → (s * t).nonempty := nonempty.image2 +@[to_additive] lemma nonempty.of_mul_left : (s * t).nonempty → s.nonempty := nonempty.of_image2_left +@[to_additive] lemma nonempty.of_mul_right : (s * t).nonempty → t.nonempty := +nonempty.of_image2_right @[simp, to_additive] lemma mul_singleton : s * {b} = (* b) '' s := image2_singleton_right @[simp, to_additive] lemma singleton_mul : {a} * t = ((*) a) '' t := image2_singleton_left +@[simp, to_additive] lemma singleton_mul_singleton : ({a} : set α) * {b} = {a * b} := +image2_singleton -@[simp, to_additive] -lemma singleton_mul_singleton : ({a} : set α) * {b} = {a * b} := image2_singleton +@[to_additive, mono] lemma mul_subset_mul : s₁ ⊆ t₁ → s₂ ⊆ t₂ → s₁ * s₂ ⊆ t₁ * t₂ := image2_subset +@[to_additive] lemma mul_subset_mul_left : t₁ ⊆ t₂ → s * t₁ ⊆ s * t₂ := image2_subset_left +@[to_additive] lemma mul_subset_mul_right : s₁ ⊆ s₂ → s₁ * t ⊆ s₂ * t := image2_subset_right +@[to_additive] lemma mul_subset_iff : s * t ⊆ u ↔ ∀ (x ∈ s) (y ∈ t), x * y ∈ u := image2_subset_iff -@[to_additive] lemma mul_subset_mul_left (h : t₁ ⊆ t₂) : s * t₁ ⊆ s * t₂ := image2_subset_left h -@[to_additive] lemma mul_subset_mul_right (h : s₁ ⊆ s₂) : s₁ * t ⊆ s₂ * t := image2_subset_right h +attribute [mono] add_subset_add @[to_additive] lemma union_mul : (s₁ ∪ s₂) * t = s₁ * t ∪ s₂ * t := image2_union_left @[to_additive] lemma mul_union : s * (t₁ ∪ t₂) = s * t₁ ∪ s * t₂ := image2_union_right +@[to_additive] lemma inter_mul_subset : (s₁ ∩ s₂) * t ⊆ s₁ * t ∩ (s₂ * t) := +image2_inter_subset_left +@[to_additive] lemma mul_inter_subset : s * (t₁ ∩ t₂) ⊆ s * t₁ ∩ (s * t₂) := +image2_inter_subset_right -@[to_additive] -lemma inter_mul_subset : (s₁ ∩ s₂) * t ⊆ s₁ * t ∩ (s₂ * t) := image2_inter_subset_left - -@[to_additive] -lemma mul_inter_subset : s * (t₁ ∩ t₂) ⊆ s * t₁ ∩ (s * t₂) := image2_inter_subset_right - -@[to_additive] -lemma Union_mul_left_image : (⋃ a ∈ s, (λ x, a * x) '' t) = s * t := Union_image_left _ - -@[to_additive] -lemma Union_mul_right_image : (⋃ a ∈ t, (λ x, x * a) '' s) = s * t := Union_image_right _ +@[to_additive] lemma Union_mul_left_image : (⋃ a ∈ s, ((*) a) '' t) = s * t := Union_image_left _ +@[to_additive] lemma Union_mul_right_image : (⋃ a ∈ t, (* a) '' s) = s * t := Union_image_right _ -@[to_additive] -lemma Union_mul (s : ι → set α) (t : set α) : (⋃ i, s i) * t = ⋃ i, s i * t := +@[to_additive] lemma Union_mul (s : ι → set α) (t : set α) : (⋃ i, s i) * t = ⋃ i, s i * t := image2_Union_left _ _ _ - -@[to_additive] -lemma mul_Union (s : set α) (t : ι → set α) : s * (⋃ i, t i) = ⋃ i, s * t i := +@[to_additive] lemma mul_Union (s : set α) (t : ι → set α) : s * (⋃ i, t i) = ⋃ i, s * t i := image2_Union_right _ _ _ @[to_additive] @@ -185,7 +174,6 @@ lemma mul_Inter₂_subset (s : set α) (t : Π i, κ i → set α) : s * (⋂ i j, t i j) ⊆ ⋂ i j, s * t i j := image2_Inter₂_subset_right _ _ _ -@[to_additive] lemma nonempty.mul : s.nonempty → t.nonempty → (s * t).nonempty := nonempty.image2 @[to_additive] lemma finite.mul : finite s → finite t → finite (s * t) := finite.image2 _ /-- Under `[has_mul M]`, the `singleton` map from `M` to `set M` as a `mul_hom`, that is, a map @@ -544,48 +532,45 @@ lemma image2_div : image2 has_div.div s t = s / t := rfl @[to_additive] lemma mem_div : a ∈ s / t ↔ ∃ x y, x ∈ s ∧ y ∈ t ∧ x / y = a := iff.rfl -@[to_additive] -lemma div_mem_div (ha : a ∈ s) (hb : b ∈ t) : a / b ∈ s / t := mem_image2_of_mem ha hb - -@[to_additive] -lemma div_subset_div (h₁ : s₁ ⊆ t₁) (h₂ : s₂ ⊆ t₂) : s₁ / s₂ ⊆ t₁ / t₂ := image2_subset h₁ h₂ +@[to_additive] lemma div_mem_div : a ∈ s → b ∈ t → a / b ∈ s / t := mem_image2_of_mem @[to_additive add_image_prod] lemma image_div_prod : (λ x : α × α, x.fst / x.snd) '' s ×ˢ t = s / t := image_prod _ @[simp, to_additive] lemma empty_div : ∅ / s = ∅ := image2_empty_left @[simp, to_additive] lemma div_empty : s / ∅ = ∅ := image2_empty_right - +@[simp, to_additive] lemma div_eq_empty : s / t = ∅ ↔ s = ∅ ∨ t = ∅ := image2_eq_empty_iff +@[simp, to_additive] lemma div_nonempty : (s / t).nonempty ↔ s.nonempty ∧ t.nonempty := +image2_nonempty_iff +@[to_additive] lemma nonempty.div : s.nonempty → t.nonempty → (s / t).nonempty := nonempty.image2 +@[to_additive] lemma nonempty.of_div_left : (s / t).nonempty → s.nonempty := nonempty.of_image2_left +@[to_additive] lemma nonempty.of_div_right : (s / t).nonempty → t.nonempty := +nonempty.of_image2_right @[simp, to_additive] lemma div_singleton : s / {b} = (/ b) '' s := image2_singleton_right @[simp, to_additive] lemma singleton_div : {a} / t = ((/) a) '' t := image2_singleton_left +@[simp, to_additive] lemma singleton_div_singleton : ({a} : set α) / {b} = {a / b} := +image2_singleton -@[simp, to_additive] -lemma singleton_div_singleton : ({a} : set α) / {b} = {a / b} := image2_singleton +@[to_additive, mono] lemma div_subset_div : s₁ ⊆ t₁ → s₂ ⊆ t₂ → s₁ / s₂ ⊆ t₁ / t₂ := image2_subset +@[to_additive] lemma div_subset_div_left : t₁ ⊆ t₂ → s / t₁ ⊆ s / t₂ := image2_subset_left +@[to_additive] lemma div_subset_div_right : s₁ ⊆ s₂ → s₁ / t ⊆ s₂ / t := image2_subset_right +@[to_additive] lemma div_subset_iff : s / t ⊆ u ↔ ∀ (x ∈ s) (y ∈ t), x / y ∈ u := image2_subset_iff -@[to_additive] lemma div_subset_div_left (h : t₁ ⊆ t₂) : s / t₁ ⊆ s / t₂ := image2_subset_left h -@[to_additive] lemma div_subset_div_right (h : s₁ ⊆ s₂) : s₁ / t ⊆ s₂ / t := image2_subset_right h +attribute [mono] sub_subset_sub @[to_additive] lemma union_div : (s₁ ∪ s₂) / t = s₁ / t ∪ s₂ / t := image2_union_left @[to_additive] lemma div_union : s / (t₁ ∪ t₂) = s / t₁ ∪ s / t₂ := image2_union_right +@[to_additive] lemma inter_div_subset : (s₁ ∩ s₂) / t ⊆ s₁ / t ∩ (s₂ / t) := +image2_inter_subset_left +@[to_additive] lemma div_inter_subset : s / (t₁ ∩ t₂) ⊆ s / t₁ ∩ (s / t₂) := +image2_inter_subset_right -@[to_additive] -lemma inter_div_subset : (s₁ ∩ s₂) / t ⊆ s₁ / t ∩ (s₂ / t) := image2_inter_subset_left - -@[to_additive] -lemma div_inter_subset : s / (t₁ ∩ t₂) ⊆ s / t₁ ∩ (s / t₂) := image2_inter_subset_right - -@[to_additive] -lemma Union_div_left_image : (⋃ a ∈ s, (λ x, a / x) '' t) = s / t := Union_image_left _ - -@[to_additive] -lemma Union_div_right_image : (⋃ a ∈ t, (λ x, x / a) '' s) = s / t := Union_image_right _ +@[to_additive] lemma Union_div_left_image : (⋃ a ∈ s, ((/) a) '' t) = s / t := Union_image_left _ +@[to_additive] lemma Union_div_right_image : (⋃ a ∈ t, (/ a) '' s) = s / t := Union_image_right _ -@[to_additive] -lemma Union_div (s : ι → set α) (t : set α) : (⋃ i, s i) / t = ⋃ i, s i / t := +@[to_additive] lemma Union_div (s : ι → set α) (t : set α) : (⋃ i, s i) / t = ⋃ i, s i / t := image2_Union_left _ _ _ - -@[to_additive] -lemma div_Union (s : set α) (t : ι → set α) : s / (⋃ i, t i) = ⋃ i, s / t i := +@[to_additive] lemma div_Union (s : set α) (t : ι → set α) : s / (⋃ i, t i) = ⋃ i, s / t i := image2_Union_right _ _ _ @[to_additive] @@ -688,47 +673,42 @@ lemma image_smul_prod : (λ x : α × β, x.fst • x.snd) '' s ×ˢ t = s • t @[to_additive] lemma mem_smul : b ∈ s • t ↔ ∃ x y, x ∈ s ∧ y ∈ t ∧ x • y = b := iff.rfl -@[to_additive] -lemma smul_mem_smul (ha : a ∈ s) (hb : b ∈ t) : a • b ∈ s • t := mem_image2_of_mem ha hb - -@[to_additive] -lemma smul_subset_smul (hs : s₁ ⊆ s₂) (ht : t₁ ⊆ t₂) : s₁ • t₁ ⊆ s₂ • t₂ := image2_subset hs ht - -@[to_additive] lemma smul_subset_iff : s • t ⊆ u ↔ ∀ (a ∈ s) (b ∈ t), a • b ∈ u := image2_subset_iff +@[to_additive] lemma smul_mem_smul : a ∈ s → b ∈ t → a • b ∈ s • t := mem_image2_of_mem @[simp, to_additive] lemma empty_smul : (∅ : set α) • t = ∅ := image2_empty_left @[simp, to_additive] lemma smul_empty : s • (∅ : set β) = ∅ := image2_empty_right - +@[simp, to_additive] lemma smul_eq_empty : s • t = ∅ ↔ s = ∅ ∨ t = ∅ := image2_eq_empty_iff +@[simp, to_additive] lemma smul_nonempty : (s • t).nonempty ↔ s.nonempty ∧ t.nonempty := +image2_nonempty_iff +@[to_additive] lemma nonempty.smul : s.nonempty → t.nonempty → (s • t).nonempty := nonempty.image2 +@[to_additive] lemma nonempty.of_smul_left : (s • t).nonempty → s.nonempty := +nonempty.of_image2_left +@[to_additive] lemma nonempty.of_smul_right : (s • t).nonempty → t.nonempty := +nonempty.of_image2_right @[simp, to_additive] lemma smul_singleton : s • {b} = (• b) '' s := image2_singleton_right @[simp, to_additive] lemma singleton_smul : ({a} : set α) • t = a • t := image2_singleton_left +@[simp, to_additive] lemma singleton_smul_singleton : ({a} : set α) • ({b} : set β) = {a • b} := +image2_singleton -@[simp, to_additive] -lemma singleton_smul_singleton : ({a} : set α) • ({b} : set β) = {a • b} := image2_singleton +@[to_additive, mono] lemma smul_subset_smul : s₁ ⊆ s₂ → t₁ ⊆ t₂ → s₁ • t₁ ⊆ s₂ • t₂ := image2_subset +@[to_additive] lemma smul_subset_smul_left : t₁ ⊆ t₂ → s • t₁ ⊆ s • t₂ := image2_subset_left +@[to_additive] lemma smul_subset_smul_right : s₁ ⊆ s₂ → s₁ • t ⊆ s₂ • t := image2_subset_right +@[to_additive] lemma smul_subset_iff : s • t ⊆ u ↔ ∀ (a ∈ s) (b ∈ t), a • b ∈ u := image2_subset_iff -@[to_additive] lemma smul_subset_smul_left (h : t₁ ⊆ t₂) : s • t₁ ⊆ s • t₂ := image2_subset_left h -@[to_additive] lemma smul_subset_smul_right (h : s₁ ⊆ s₂) : s₁ • t ⊆ s₂ • t := image2_subset_right h +attribute [mono] vadd_subset_vadd @[to_additive] lemma union_smul : (s₁ ∪ s₂) • t = s₁ • t ∪ s₂ • t := image2_union_left @[to_additive] lemma smul_union : s • (t₁ ∪ t₂) = s • t₁ ∪ s • t₂ := image2_union_right +@[to_additive] lemma inter_smul_subset : (s₁ ∩ s₂) • t ⊆ s₁ • t ∩ s₂ • t := image2_inter_subset_left +@[to_additive] lemma smul_inter_subset : s • (t₁ ∩ t₂) ⊆ s • t₁ ∩ s • t₂ := +image2_inter_subset_right -@[to_additive] -lemma inter_smul_subset : (s₁ ∩ s₂) • t ⊆ s₁ • t ∩ s₂ • t := image2_inter_subset_left - -@[to_additive] -lemma smul_inter_subset : s • (t₁ ∩ t₂) ⊆ s • t₁ ∩ s • t₂ := image2_inter_subset_right - -@[to_additive] -lemma Union_smul_left_image : (⋃ a ∈ s, a • t) = s • t := Union_image_left _ - -@[to_additive] -lemma Union_smul_right_image : (⋃ a ∈ t, (λ x, x • a) '' s) = s • t := Union_image_right _ +@[to_additive] lemma Union_smul_left_image : (⋃ a ∈ s, a • t) = s • t := Union_image_left _ +@[to_additive] lemma Union_smul_right_image : (⋃ a ∈ t, (• a) '' s) = s • t := Union_image_right _ -@[to_additive] -lemma Union_smul (s : ι → set α) (t : set β) : (⋃ i, s i) • t = ⋃ i, s i • t := +@[to_additive] lemma Union_smul (s : ι → set α) (t : set β) : (⋃ i, s i) • t = ⋃ i, s i • t := image2_Union_left _ _ _ - -@[to_additive] -lemma smul_Union (s : set α) (t : ι → set β) : s • (⋃ i, t i) = ⋃ i, s • t i := +@[to_additive] lemma smul_Union (s : set α) (t : ι → set β) : s • (⋃ i, t i) = ⋃ i, s • t i := image2_Union_right _ _ _ @[to_additive] @@ -757,7 +737,6 @@ lemma smul_Inter₂_subset (s : set α) (t : Π i, κ i → set β) : s • (⋂ i j, t i j) ⊆ ⋂ i j, s • t i j := image2_Inter₂_subset_right _ _ _ -@[to_additive] lemma nonempty.smul : s.nonempty → t.nonempty → (s • t).nonempty := nonempty.image2 @[to_additive] lemma finite.smul : finite s → finite t → finite (s • t) := finite.image2 _ end has_scalar @@ -769,13 +748,11 @@ variables {ι : Sort*} {κ : ι → Sort*} [has_scalar α β] {s t t₁ t₂ : s @[to_additive] lemma mem_smul_set : x ∈ a • t ↔ ∃ y, y ∈ t ∧ a • y = x := iff.rfl -@[to_additive] lemma smul_mem_smul_set (hy : y ∈ t) : a • y ∈ a • t := ⟨y, hy, rfl⟩ - -@[to_additive] -lemma mem_smul_of_mem {s : set α} (ha : a ∈ s) (hb : b ∈ t) : a • b ∈ s • t := -mem_image2_of_mem ha hb +@[to_additive] lemma smul_mem_smul_set : b ∈ s → a • b ∈ a • s := mem_image_of_mem _ @[simp, to_additive] lemma smul_set_empty : a • (∅ : set β) = ∅ := image_empty _ +@[simp, to_additive] lemma smul_set_eq_empty : a • s = ∅ ↔ s = ∅ := image_eq_empty +@[simp, to_additive] lemma smul_set_nonempty : (a • s).nonempty ↔ s.nonempty := nonempty_image_iff @[simp, to_additive] lemma smul_set_singleton : a • ({b} : set β) = {a • b} := image_singleton @@ -905,7 +882,8 @@ localized "attribute [instance] set.distrib_mul_action_set set.mul_distrib_mul_a end smul section vsub -variables {ι : Sort*} {κ : ι → Sort*} [has_vsub α β] {s s₁ s₂ t t₁ t₂ : set β} {a : α} {b c : β} +variables {ι : Sort*} {κ : ι → Sort*} [has_vsub α β] {s s₁ s₂ t t₁ t₂ : set β} {u : set α} {a : α} + {b c : β} include α instance has_vsub : has_vsub (set α) (set β) := ⟨image2 (-ᵥ)⟩ @@ -918,24 +896,26 @@ lemma mem_vsub : a ∈ s -ᵥ t ↔ ∃ x y, x ∈ s ∧ y ∈ t ∧ x -ᵥ y = lemma vsub_mem_vsub (hb : b ∈ s) (hc : c ∈ t) : b -ᵥ c ∈ s -ᵥ t := mem_image2_of_mem hb hc -lemma vsub_subset_vsub (hs : s₁ ⊆ s₂) (ht : t₁ ⊆ t₂) : s₁ -ᵥ t₁ ⊆ s₂ -ᵥ t₂ := image2_subset hs ht - -lemma vsub_subset_iff {u : set α} : s -ᵥ t ⊆ u ↔ ∀ (x ∈ s) (y ∈ t), x -ᵥ y ∈ u := image2_subset_iff - @[simp] lemma empty_vsub (t : set β) : ∅ -ᵥ t = ∅ := image2_empty_left @[simp] lemma vsub_empty (s : set β) : s -ᵥ ∅ = ∅ := image2_empty_right - +@[simp] lemma vsub_eq_empty : s -ᵥ t = ∅ ↔ s = ∅ ∨ t = ∅ := image2_eq_empty_iff +@[simp] lemma vsub_nonempty : (s -ᵥ t : set α).nonempty ↔ s.nonempty ∧ t.nonempty := +image2_nonempty_iff +lemma nonempty.vsub : s.nonempty → t.nonempty → (s -ᵥ t : set α).nonempty := nonempty.image2 +lemma nonempty.of_vsub_left : (s -ᵥ t :set α).nonempty → s.nonempty := nonempty.of_image2_left +lemma nonempty.of_vsub_right : (s -ᵥ t : set α).nonempty → t.nonempty := nonempty.of_image2_right @[simp] lemma vsub_singleton (s : set β) (b : β) : s -ᵥ {b} = (-ᵥ b) '' s := image2_singleton_right @[simp] lemma singleton_vsub (t : set β) (b : β) : {b} -ᵥ t = ((-ᵥ) b) '' t := image2_singleton_left - @[simp] lemma singleton_vsub_singleton : ({b} : set β) -ᵥ {c} = {b -ᵥ c} := image2_singleton -lemma vsub_subset_vsub_left (h : t₁ ⊆ t₂) : s -ᵥ t₁ ⊆ s -ᵥ t₂ := image2_subset_left h -lemma vsub_subset_vsub_right (h : s₁ ⊆ s₂) : s₁ -ᵥ t ⊆ s₂ -ᵥ t := image2_subset_right h +@[mono] lemma vsub_subset_vsub : s₁ ⊆ s₂ → t₁ ⊆ t₂ → s₁ -ᵥ t₁ ⊆ s₂ -ᵥ t₂ := image2_subset +lemma vsub_subset_vsub_left : t₁ ⊆ t₂ → s -ᵥ t₁ ⊆ s -ᵥ t₂ := image2_subset_left +lemma vsub_subset_vsub_right : s₁ ⊆ s₂ → s₁ -ᵥ t ⊆ s₂ -ᵥ t := image2_subset_right +lemma vsub_subset_iff : s -ᵥ t ⊆ u ↔ ∀ (x ∈ s) (y ∈ t), x -ᵥ y ∈ u := image2_subset_iff +lemma vsub_self_mono (h : s ⊆ t) : s -ᵥ s ⊆ t -ᵥ t := vsub_subset_vsub h h lemma union_vsub : (s₁ ∪ s₂) -ᵥ t = s₁ -ᵥ t ∪ (s₂ -ᵥ t) := image2_union_left lemma vsub_union : s -ᵥ (t₁ ∪ t₂) = s -ᵥ t₁ ∪ (s -ᵥ t₂) := image2_union_right - lemma inter_vsub_subset : s₁ ∩ s₂ -ᵥ t ⊆ (s₁ -ᵥ t) ∩ (s₂ -ᵥ t) := image2_inter_subset_left lemma vsub_inter_subset : s -ᵥ t₁ ∩ t₂ ⊆ (s -ᵥ t₁) ∩ (s -ᵥ t₂) := image2_inter_subset_right @@ -944,7 +924,6 @@ lemma Union_vsub_right_image : (⋃ a ∈ t, (-ᵥ a) '' s) = s -ᵥ t := Union_ lemma Union_vsub (s : ι → set β) (t : set β) : (⋃ i, s i) -ᵥ t = ⋃ i, s i -ᵥ t := image2_Union_left _ _ _ - lemma vsub_Union (s : set β) (t : ι → set β) : s -ᵥ (⋃ i, t i) = ⋃ i, s -ᵥ t i := image2_Union_right _ _ _ @@ -968,11 +947,8 @@ lemma vsub_Inter₂_subset (s : set β) (t : Π i, κ i → set β) : s -ᵥ (⋂ i j, t i j) ⊆ ⋂ i j, s -ᵥ t i j := image2_Inter₂_subset_right _ _ _ -lemma nonempty.vsub : s.nonempty → t.nonempty → (s -ᵥ t : set α).nonempty := nonempty.image2 lemma finite.vsub (hs : finite s) (ht : finite t) : finite (s -ᵥ t) := hs.image2 _ ht -lemma vsub_self_mono (h : s ⊆ t) : s -ᵥ s ⊆ t -ᵥ t := vsub_subset_vsub h h - end vsub open_locale pointwise diff --git a/src/order/filter/n_ary.lean b/src/order/filter/n_ary.lean index 7feb6c067c9a0..153dcbd3b092d 100644 --- a/src/order/filter/n_ary.lean +++ b/src/order/filter/n_ary.lean @@ -114,6 +114,9 @@ by { simp_rw ne_bot_iff, exact map₂_eq_bot_iff.not.trans not_or_distrib } lemma ne_bot.map₂ (hf : f.ne_bot) (hg : g.ne_bot) : (map₂ m f g).ne_bot := map₂_ne_bot_iff.2 ⟨hf, hg⟩ +lemma ne_bot.of_map₂_left (h : (map₂ m f g).ne_bot) : f.ne_bot := (map₂_ne_bot_iff.1 h).1 +lemma ne_bot.of_map₂_right (h : (map₂ m f g).ne_bot) : g.ne_bot := (map₂_ne_bot_iff.1 h).2 + lemma map₂_sup_left : map₂ m (f₁ ⊔ f₂) g = map₂ m f₁ g ⊔ map₂ m f₂ g := begin ext u, diff --git a/src/order/filter/pointwise.lean b/src/order/filter/pointwise.lean index 490e514eb3054..b8517656fcbc0 100644 --- a/src/order/filter/pointwise.lean +++ b/src/order/filter/pointwise.lean @@ -5,6 +5,7 @@ Authors: Zhouhang Zhou, Yaël Dillies -/ import data.set.pointwise import order.filter.n_ary +import order.filter.ultrafilter /-! # Pointwise operations on filters @@ -71,8 +72,8 @@ localized "attribute [instance] filter.has_one filter.has_zero" in pointwise @[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 pure_one : pure 1 = (1 : filter α) := (principal_singleton _).symm @[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 tendsto_one {a : filter β} {f : β → α} : @@ -92,7 +93,7 @@ end one /-! ### Filter addition/multiplication -/ section mul -variables [has_mul α] [has_mul β] {f f₁ f₂ g g₁ g₂ h : filter α} {s t : set α} +variables [has_mul α] [has_mul β] {f f₁ f₂ g g₁ g₂ h : filter α} {s t : set α} {a b : α} /-- 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`."] @@ -108,6 +109,11 @@ localized "attribute [instance] filter.has_mul filter.has_add" in pointwise @[simp, to_additive] lemma mul_eq_bot_iff : f * g = ⊥ ↔ f = ⊥ ∨ g = ⊥ := map₂_eq_bot_iff @[simp, to_additive] lemma mul_ne_bot_iff : (f * g).ne_bot ↔ f.ne_bot ∧ g.ne_bot := map₂_ne_bot_iff @[to_additive] lemma ne_bot.mul : ne_bot f → ne_bot g → ne_bot (f * g) := ne_bot.map₂ +@[to_additive] lemma ne_bot.of_mul_left : (f * g).ne_bot → f.ne_bot := ne_bot.of_map₂_left +@[to_additive] lemma ne_bot.of_mul_right : (f * g).ne_bot → g.ne_bot := ne_bot.of_map₂_right +@[simp, to_additive] lemma pure_mul : pure a * g = g.map ((*) a) := map₂_pure_left +@[simp, to_additive] lemma mul_pure : f * pure b = f.map (* b) := map₂_pure_right +@[simp, to_additive] lemma pure_mul_pure : (pure a : filter α) * pure b = pure (a * b) := map₂_pure @[simp, to_additive] lemma le_mul_iff : h ≤ f * g ↔ ∀ ⦃s⦄, s ∈ f → ∀ ⦃t⦄, t ∈ g → s * t ∈ h := le_map₂_iff @@ -190,7 +196,7 @@ end map /-! ### Filter negation/inversion -/ section has_inv -variables [has_inv α] {f g : filter α} {s : set α} +variables [has_inv α] {f g : filter α} {s : set α} {a : α} /-- The inverse of a filter is the pointwise preimage under `⁻¹` of its sets. -/ @[to_additive "The negation of a filter is the pointwise preimage under `-` of its sets."] @@ -199,6 +205,8 @@ instance : has_inv (filter α) := ⟨map has_inv.inv⟩ @[simp, to_additive] protected lemma map_inv : f.map has_inv.inv = f⁻¹ := rfl @[to_additive] lemma mem_inv : s ∈ f⁻¹ ↔ has_inv.inv ⁻¹' s ∈ f := iff.rfl @[to_additive] protected lemma inv_le_inv (hf : f ≤ g) : f⁻¹ ≤ g⁻¹ := map_mono hf +@[simp, to_additive] lemma inv_pure : (pure a : filter α)⁻¹ = pure a⁻¹ := rfl +@[simp, to_additive] lemma inv_eq_bot_iff : f⁻¹ = ⊥ ↔ f = ⊥ := map_eq_bot_iff @[simp, to_additive] lemma ne_bot_inv_iff : f⁻¹.ne_bot ↔ ne_bot f := map_ne_bot_iff _ @[to_additive] lemma ne_bot.inv : f.ne_bot → f⁻¹.ne_bot := λ h, h.map _ @@ -234,7 +242,7 @@ end group /-! ### Filter subtraction/division -/ section div -variables [has_div α] {f f₁ f₂ g g₁ g₂ h : filter α} {s t : set α} +variables [has_div α] {f f₁ f₂ g g₁ g₂ h : filter α} {s t : set α} {a b : α} /-- 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`."] @@ -250,12 +258,17 @@ localized "attribute [instance] filter.has_div filter.has_sub" in pointwise @[simp, to_additive] lemma div_eq_bot_iff : f / g = ⊥ ↔ f = ⊥ ∨ g = ⊥ := map₂_eq_bot_iff @[simp, to_additive] lemma div_ne_bot_iff : (f / g).ne_bot ↔ f.ne_bot ∧ g.ne_bot := map₂_ne_bot_iff @[to_additive] lemma ne_bot.div : ne_bot f → ne_bot g → ne_bot (f / g) := ne_bot.map₂ -@[simp, to_additive] protected lemma le_div_iff : - h ≤ f / g ↔ ∀ ⦃s⦄, s ∈ f → ∀ ⦃t⦄, t ∈ g → s / t ∈ h := -le_map₂_iff +@[to_additive] lemma ne_bot.of_div_left : (f / g).ne_bot → f.ne_bot := ne_bot.of_map₂_left +@[to_additive] lemma ne_bot.of_div_right : (f / g).ne_bot → g.ne_bot := ne_bot.of_map₂_right +@[simp, to_additive] lemma pure_div : pure a / g = g.map ((/) a) := map₂_pure_left +@[simp, to_additive] lemma div_pure : f / pure b = f.map (/ b) := map₂_pure_right +@[simp, to_additive] lemma pure_div_pure : (pure a : filter α) / pure b = pure (a / b) := map₂_pure @[to_additive] protected lemma div_le_div : f₁ ≤ f₂ → g₁ ≤ g₂ → f₁ / g₁ ≤ f₂ / g₂ := map₂_mono @[to_additive] protected lemma div_le_div_left : g₁ ≤ g₂ → f / g₁ ≤ f / g₂ := map₂_mono_left @[to_additive] protected lemma div_le_div_right : f₁ ≤ f₂ → f₁ / g ≤ f₂ / g := map₂_mono_right +@[simp, to_additive] protected lemma le_div_iff : + h ≤ f / g ↔ ∀ ⦃s⦄, s ∈ f → ∀ ⦃t⦄, t ∈ g → s / t ∈ h := +le_map₂_iff @[to_additive] instance covariant_div : covariant_class (filter α) (filter α) (/) (≤) := ⟨λ f g h, map₂_mono_left⟩ @@ -303,6 +316,7 @@ localized "attribute [instance] filter.div_inv_monoid filter.sub_neg_monoid filt 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₂ (•)⟩ @@ -314,11 +328,17 @@ variables [has_scalar α β] {f f₁ f₂ : filter α} {g g₁ g₂ h : filter @[simp, to_additive] lemma smul_eq_bot_iff : f • g = ⊥ ↔ f = ⊥ ∨ g = ⊥ := map₂_eq_bot_iff @[simp, to_additive] lemma smul_ne_bot_iff : (f • g).ne_bot ↔ f.ne_bot ∧ g.ne_bot := map₂_ne_bot_iff @[to_additive] lemma ne_bot.smul : ne_bot f → ne_bot g → ne_bot (f • g) := ne_bot.map₂ -@[simp, to_additive] lemma le_smul_iff : h ≤ f • g ↔ ∀ ⦃s⦄, s ∈ f → ∀ ⦃t⦄, t ∈ g → s • t ∈ h := -le_map₂_iff +@[to_additive] lemma ne_bot.of_smul_left : (f • g).ne_bot → f.ne_bot := ne_bot.of_map₂_left +@[to_additive] lemma ne_bot.of_smul_right : (f • g).ne_bot → g.ne_bot := ne_bot.of_map₂_right +@[simp, to_additive] lemma pure_smul : (pure a : filter α) • g = g.map ((•) a) := map₂_pure_left +@[simp, to_additive] lemma smul_pure : f • pure b = f.map (• b) := map₂_pure_right +@[simp, to_additive] lemma pure_smul_pure : + (pure a : filter α) • (pure b : filter β) = pure (a • b) := map₂_pure @[to_additive] lemma smul_le_smul : f₁ ≤ f₂ → g₁ ≤ g₂ → f₁ • g₁ ≤ f₂ • g₂ := map₂_mono @[to_additive] lemma smul_le_smul_left : g₁ ≤ g₂ → f • g₁ ≤ f • g₂ := map₂_mono_left @[to_additive] lemma smul_le_smul_right : f₁ ≤ f₂ → f₁ • g ≤ f₂ • g := map₂_mono_right +@[simp, to_additive] lemma le_smul_iff : h ≤ f • g ↔ ∀ ⦃s⦄, s ∈ f → ∀ ⦃t⦄, t ∈ g → s • t ∈ h := +le_map₂_iff @[to_additive] instance covariant_smul : covariant_class (filter α) (filter β) (•) (≤) := ⟨λ f g h, map₂_mono_left⟩ @@ -333,7 +353,7 @@ instance [monoid α] [mul_action α β] : mul_action (filter α) (filter β) := /-! ### Scalar subtraction of filters -/ section vsub -variables [has_vsub α β] {f f₁ f₂ g g₁ g₂ : filter β} {h : filter α} {s t : set β} +variables [has_vsub α β] {f f₁ f₂ g g₁ g₂ : filter β} {h : filter α} {s t : set β} {a b : β} include α /-- The filter `f -ᵥ g` is generated by `{s -ᵥ t | s ∈ f, t ∈ g}` in locale `pointwise`. -/ @@ -349,10 +369,16 @@ lemma vsub_mem_vsub : s ∈ f → t ∈ g → s -ᵥ t ∈ f -ᵥ g := image2_m @[simp] lemma vsub_eq_bot_iff : f -ᵥ g = ⊥ ↔ f = ⊥ ∨ g = ⊥ := map₂_eq_bot_iff @[simp] lemma vsub_ne_bot_iff : (f -ᵥ g : filter α).ne_bot ↔ f.ne_bot ∧ g.ne_bot := map₂_ne_bot_iff lemma ne_bot.vsub : ne_bot f → ne_bot g → ne_bot (f -ᵥ g) := ne_bot.map₂ -@[simp] lemma le_vsub_iff : h ≤ f -ᵥ g ↔ ∀ ⦃s⦄, s ∈ f → ∀ ⦃t⦄, t ∈ g → s -ᵥ t ∈ h := le_map₂_iff +lemma ne_bot.of_vsub_left : (f -ᵥ g : filter α).ne_bot → f.ne_bot := ne_bot.of_map₂_left +lemma ne_bot.of_vsub_right : (f -ᵥ g : filter α).ne_bot → g.ne_bot := ne_bot.of_map₂_right +@[simp] lemma pure_vsub : (pure a : filter β) -ᵥ g = g.map ((-ᵥ) a) := map₂_pure_left +@[simp] lemma vsub_pure : f -ᵥ pure b = f.map (-ᵥ b) := map₂_pure_right +@[simp] lemma pure_vsub_pure : (pure a : filter β) -ᵥ pure b = (pure (a -ᵥ b) : filter α) := +map₂_pure lemma vsub_le_vsub : f₁ ≤ f₂ → g₁ ≤ g₂ → f₁ -ᵥ g₁ ≤ f₂ -ᵥ g₂ := map₂_mono lemma vsub_le_vsub_left : g₁ ≤ g₂ → f -ᵥ g₁ ≤ f -ᵥ g₂ := map₂_mono_left lemma vsub_le_vsub_right : f₁ ≤ f₂ → f₁ -ᵥ g ≤ f₂ -ᵥ g := map₂_mono_right +@[simp] lemma le_vsub_iff : h ≤ f -ᵥ g ↔ ∀ ⦃s⦄, s ∈ f → ∀ ⦃t⦄, t ∈ g → s -ᵥ t ∈ h := le_map₂_iff end vsub diff --git a/src/topology/algebra/filter_basis.lean b/src/topology/algebra/filter_basis.lean index fe0463d536655..9df8133636bf2 100644 --- a/src/topology/algebra/filter_basis.lean +++ b/src/topology/algebra/filter_basis.lean @@ -364,7 +364,7 @@ begin rcases B.smul U_in with ⟨V, V_in, W, W_in, H⟩, apply mem_of_superset (prod_mem_prod V_in $ B'.mem_nhds_zero W_in), rintros ⟨v, w⟩ ⟨v_in : v ∈ V, w_in : w ∈ W⟩, - exact H (set.mem_smul_of_mem v_in w_in) }, + exact H (set.smul_mem_smul v_in w_in) }, { intro m₀, rw basis.tendsto_right_iff, intros U U_in, From da06587c1c602c9e729dedbdc26811c70d6a1eb3 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Thu, 5 May 2022 11:55:58 +0000 Subject: [PATCH 142/153] feat(linear_algebra): A-linear maps between finite dimensional vector spaces over k are finite dimensional (#13934) Co-authored-by: Scott Morrison --- src/algebra/module/algebra.lean | 32 +++++++++++++++++++++++++++ src/linear_algebra/matrix/to_lin.lean | 16 +++++++++++++- 2 files changed, 47 insertions(+), 1 deletion(-) create mode 100644 src/algebra/module/algebra.lean diff --git a/src/algebra/module/algebra.lean b/src/algebra/module/algebra.lean new file mode 100644 index 0000000000000..129e92cffef04 --- /dev/null +++ b/src/algebra/module/algebra.lean @@ -0,0 +1,32 @@ +/- +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.module.basic +import algebra.algebra.basic + +/-! +# Additional facts about modules over algebras. +-/ + +namespace linear_map + +section restrict_scalars + +variables (k : Type*) [comm_semiring k] (A : Type*) [semiring A] [algebra k A] +variables (M : Type*) [add_comm_monoid M] [module k M] [module A M] [is_scalar_tower k A M] +variables (N : Type*) [add_comm_monoid N] [module k N] [module A N] [is_scalar_tower k A N] + +/-- +Restriction of scalars for linear maps between modules over a `k`-algebra is itself `k`-linear. +-/ +@[simps] +def restrict_scalars_linear_map : (M →ₗ[A] N) →ₗ[k] (M →ₗ[k] N) := +{ to_fun := linear_map.restrict_scalars k, + map_add' := by tidy, + map_smul' := by tidy, } + +end restrict_scalars + +end linear_map diff --git a/src/linear_algebra/matrix/to_lin.lean b/src/linear_algebra/matrix/to_lin.lean index 7b2dfbfbc4e32..bfceaddb01a3a 100644 --- a/src/linear_algebra/matrix/to_lin.lean +++ b/src/linear_algebra/matrix/to_lin.lean @@ -7,6 +7,7 @@ import data.matrix.block import linear_algebra.matrix.finite_dimensional import linear_algebra.std_basis import ring_theory.algebra_tower +import algebra.module.algebra /-! # Linear maps and matrices @@ -593,10 +594,23 @@ variables {K : Type*} [field K] variables {V : Type*} [add_comm_group V] [module K V] [finite_dimensional K V] variables {W : Type*} [add_comm_group W] [module K W] [finite_dimensional K W] -instance : finite_dimensional K (V →ₗ[K] W) := +instance finite_dimensional : finite_dimensional K (V →ₗ[K] W) := linear_equiv.finite_dimensional (linear_map.to_matrix (basis.of_vector_space K V) (basis.of_vector_space K W)).symm +section + +variables {A : Type*} [ring A] [algebra K A] [module A V] [is_scalar_tower K A V] + [module A W] [is_scalar_tower K A W] + +/-- Linear maps over a `k`-algebra are finite dimensional (over `k`) if both the source and +target are, since they form a subspace of all `k`-linear maps. -/ +instance finite_dimensional' : finite_dimensional K (V →ₗ[A] W) := +finite_dimensional.of_injective (restrict_scalars_linear_map K A V W) + (restrict_scalars_injective _) + +end + /-- The dimension of the space of linear transformations is the product of the dimensions of the domain and codomain. From 057e0283c5a877ddf8286ee8168c7fb17703349e Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Thu, 5 May 2022 11:55:59 +0000 Subject: [PATCH 143/153] feat(linear_algebra/finite_dimensional): surjective_of_nonzero_of_finrank_eq_one (#13961) Co-authored-by: Scott Morrison --- src/linear_algebra/finite_dimensional.lean | 16 ++++++++++++++++ 1 file changed, 16 insertions(+) diff --git a/src/linear_algebra/finite_dimensional.lean b/src/linear_algebra/finite_dimensional.lean index 3fda0e8713624..de89d6e85bdc2 100644 --- a/src/linear_algebra/finite_dimensional.lean +++ b/src/linear_algebra/finite_dimensional.lean @@ -1642,6 +1642,22 @@ begin exact finrank_le_one v p, } end +-- We use the `linear_map.compatible_smul` typeclass here, to encompass two situations: +-- * `A = K` +-- * `[field K] [algebra K A] [is_scalar_tower K A V] [is_scalar_tower K A W]` +lemma surjective_of_nonzero_of_finrank_eq_one + {K : Type*} [division_ring K] {A : Type*} [semiring A] + [module K V] [module A V] + {W : Type*} [add_comm_group W] [module K W] [module A W] [linear_map.compatible_smul V W K A] + (h : finrank K W = 1) {f : V →ₗ[A] W} (w : f ≠ 0) : surjective f := +begin + change surjective (f.restrict_scalars K), + obtain ⟨v, n⟩ := fun_like.ne_iff.mp w, + intro z, + obtain ⟨c, rfl⟩ := (finrank_eq_one_iff_of_nonzero' (f v) n).mp h z, + exact ⟨c • v, by simp⟩, +end + end finrank_eq_one section subalgebra_dim From 8e0ab168514b96107f000bf5dcf906508b19a1e0 Mon Sep 17 00:00:00 2001 From: Eric Rodriguez Date: Thu, 5 May 2022 13:10:58 +0000 Subject: [PATCH 144/153] =?UTF-8?q?feat(polynomial/cyclotomic/basic):=20?= =?UTF-8?q?=C9=B8=5Fp=E2=81=B1=20irreducible=20=E2=86=92=20=C9=B8=5Fp?= =?UTF-8?q?=CA=B2=20irreducible=20for=20j=20=E2=89=A4=20i=20(#13952)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Eric Rodriguez <37984851+ericrbg@users.noreply.github.com> --- .../cyclotomic/discriminant.lean | 5 +- .../cyclotomic/primitive_roots.lean | 47 +++++++++---------- .../polynomial/cyclotomic/basic.lean | 30 +++++++----- 3 files changed, 42 insertions(+), 40 deletions(-) diff --git a/src/number_theory/cyclotomic/discriminant.lean b/src/number_theory/cyclotomic/discriminant.lean index 6632b37d9aaa2..ef6c2e982b731 100644 --- a/src/number_theory/cyclotomic/discriminant.lean +++ b/src/number_theory/cyclotomic/discriminant.lean @@ -66,7 +66,6 @@ lemma discr_prime_pow_ne_two [is_cyclotomic_extension {p ^ (k + 1)} K L] [hp : f discr K (hζ.power_basis K).basis = (-1) ^ (((p ^ (k + 1) : ℕ).totient) / 2) * p ^ ((p : ℕ) ^ k * ((p - 1) * (k + 1) - 1)) := begin - have hirr₁ := cyclotomic_irreducible_of_irreducible_pow hp.1 k.succ_ne_zero hirr, haveI : ne_zero ((↑(p ^ (k + 1)) : ℕ) : K), { refine ⟨λ hzero, _⟩, rw [pnat.pow_coe] at hzero, @@ -116,8 +115,8 @@ begin replace H := congr_arg (algebra.norm K) H, have hnorm : (norm K) (ζ ^ (p : ℕ) ^ k - 1) = p ^ ((p : ℕ) ^ k), { by_cases hp : p = 2, - { exact hζ.pow_sub_one_norm_prime_pow_of_one_le hirr (by simpa using hirr₁) rfl.le (hp2 hp) }, - { exact hζ.pow_sub_one_norm_prime_ne_two hirr (by simpa using hirr₁) rfl.le hp } }, + { exact hζ.pow_sub_one_norm_prime_pow_of_one_le hirr rfl.le (hp2 hp) }, + { exact hζ.pow_sub_one_norm_prime_ne_two hirr rfl.le hp } }, rw [monoid_hom.map_mul, hnorm, monoid_hom.map_mul, ← map_nat_cast (algebra_map K L), algebra.norm_algebra_map, finrank _ hirr, pnat.pow_coe, totient_prime_pow hp.out (succ_pos k), nat.sub_one, nat.pred_succ, ← hζ.minpoly_eq_cyclotomic_of_irreducible hirr, map_pow, diff --git a/src/number_theory/cyclotomic/primitive_roots.lean b/src/number_theory/cyclotomic/primitive_roots.lean index ce0960555630a..a5db1f513fb72 100644 --- a/src/number_theory/cyclotomic/primitive_roots.lean +++ b/src/number_theory/cyclotomic/primitive_roots.lean @@ -37,11 +37,10 @@ in the implementation details section. * `is_primitive_root.sub_one_norm_eq_eval_cyclotomic`: if `irreducible (cyclotomic n K)` (in particular for `K = ℚ`), then the norm of `ζ - 1` is `eval 1 (cyclotomic n ℤ)`, for a primitive root ζ. We also prove the analogous of this result for `zeta`. -* `is_primitive_root.pow_prime_pow_sub_one_norm` : if - `irreducible (cyclotomic (p ^ (k + 1)) K)` and `irreducible (cyclotomic (p ^ (k - s + 1)) K))` - (in particular for `K = ℚ`) and `p` is a prime, then the norm of `ζ ^ (p ^ s) - 1` is - `p ^ (p ^ s)` `p ^ (k - s + 1) ≠ 2`. See the following lemmas for similar results. We also prove - the analogous of this result for `zeta`. +* `is_primitive_root.pow_sub_one_norm_prime_pow_ne_two` : if + `irreducible (cyclotomic (p ^ (k + 1)) K)` (in particular for `K = ℚ`) and `p` is a prime, + then the norm of `ζ ^ (p ^ s) - 1` is `p ^ (p ^ s)` `p ^ (k - s + 1) ≠ 2`. See the following + lemmas for similar results. We also prove the analogous of this result for `zeta`. * `is_primitive_root.sub_one_norm_prime_ne_two` : if `irreducible (cyclotomic (p ^ (k + 1)) K)` (in particular for `K = ℚ`) and `p` is an odd prime, then the norm of `ζ - 1` is `p`. We also prove the analogous of this result for `zeta`. @@ -285,15 +284,13 @@ omit hζ local attribute [instance] is_cyclotomic_extension.finite_dimensional local attribute [instance] is_cyclotomic_extension.is_galois -/-- If `irreducible (cyclotomic (p ^ (k + 1)) K)` and -`irreducible (cyclotomic (p ^ (k - s + 1)) K))` (in particular for `K = ℚ`) and `p` is a prime, +/-- If `irreducible (cyclotomic (p ^ (k + 1)) K)` (in particular for `K = ℚ`) and `p` is a prime, then the norm of `ζ ^ (p ^ s) - 1` is `p ^ (p ^ s)` if `p ^ (k - s + 1) ≠ 2`. See the next lemmas for similar results. -/ lemma pow_sub_one_norm_prime_pow_ne_two [ne_zero ((p : ℕ) : K)] {k s : ℕ} (hζ : is_primitive_root ζ ↑(p ^ (k + 1))) [hpri : fact (p : ℕ).prime] [is_cyclotomic_extension {p ^ (k + 1)} K L] - (hirr : irreducible (cyclotomic (↑(p ^ (k + 1)) : ℕ) K)) - (hirr₁ : irreducible (cyclotomic (↑(p ^ (k - s + 1)) : ℕ) K)) (hs : s ≤ k) + (hirr : irreducible (cyclotomic (↑(p ^ (k + 1)) : ℕ) K)) (hs : s ≤ k) (htwo : p ^ (k - s + 1) ≠ 2) : norm K (ζ ^ ((p : ℕ) ^ s) - 1) = p ^ ((p : ℕ) ^ s) := begin haveI : ne_zero ((↑(p ^ (k + 1)) : ℕ) : K), @@ -305,6 +302,9 @@ begin rw [pnat.pow_coe] at hzero, simpa [ne_zero.ne ((p : ℕ) : K)] using hzero }, + have hirr₁ : irreducible (cyclotomic (p ^ (k - s + 1)) K) := + cyclotomic_irreducible_pow_of_irreducible_pow hpri.1 (by linarith) hirr, + rw ←pnat.pow_coe at hirr₁, let η := ζ ^ ((p : ℕ) ^ s) - 1, let η₁ : K⟮η⟯ := intermediate_field.adjoin_simple.gen K η, have hη : is_primitive_root (η + 1) (p ^ (k + 1 - s)), @@ -354,17 +354,15 @@ begin all_goals { apply_instance } end -/-- If `irreducible (cyclotomic (p ^ (k + 1)) K)` and -`irreducible (cyclotomic (p ^ (k - s + 1)) K))` (in particular for `K = ℚ`) and `p` is a prime, +/-- If `irreducible (cyclotomic (p ^ (k + 1)) K)` (in particular for `K = ℚ`) and `p` is a prime, then the norm of `ζ ^ (p ^ s) - 1` is `p ^ (p ^ s)` if `p ≠ 2`. -/ lemma pow_sub_one_norm_prime_ne_two [ne_zero ((p : ℕ) : K)] {k : ℕ} (hζ : is_primitive_root ζ ↑(p ^ (k + 1))) [hpri : fact (p : ℕ).prime] [is_cyclotomic_extension {p ^ (k + 1)} K L] - (hirr : irreducible (cyclotomic (↑(p ^ (k + 1)) : ℕ) K)) {s : ℕ} - (hirr₁ : irreducible (cyclotomic (↑(p ^ (k - s + 1)) : ℕ) K)) (hs : s ≤ k) + (hirr : irreducible (cyclotomic (↑(p ^ (k + 1)) : ℕ) K)) {s : ℕ} (hs : s ≤ k) (hodd : p ≠ 2) : norm K (ζ ^ ((p : ℕ) ^ s) - 1) = p ^ ((p : ℕ) ^ s) := begin - refine hζ.pow_sub_one_norm_prime_pow_ne_two hirr hirr₁ hs (λ h, _), + refine hζ.pow_sub_one_norm_prime_pow_ne_two hirr hs (λ h, _), rw [← pnat.coe_inj, pnat.coe_bit0, pnat.one_coe, pnat.pow_coe, ← pow_one 2] at h, replace h := eq_of_prime_pow_eq (prime_iff.1 hpri.out) (prime_iff.1 nat.prime_two) ((k - s).succ_pos) h, @@ -379,7 +377,7 @@ lemma sub_one_norm_prime_ne_two [ne_zero ((p : ℕ) : K)] {k : ℕ} [is_cyclotomic_extension {p ^ (k + 1)} K L] (hirr : irreducible (cyclotomic (↑(p ^ (k + 1)) : ℕ) K)) (h : p ≠ 2) : norm K (ζ - 1) = p := -by simpa using hζ.pow_sub_one_norm_prime_ne_two hirr (by exact hirr) (zero_le k) h +by simpa using hζ.pow_sub_one_norm_prime_ne_two hirr k.zero_le h /-- If `irreducible (cyclotomic p K)` (in particular for `K = ℚ`) and `p` is an odd prime, then the norm of `ζ - 1` is `p`. -/ @@ -441,14 +439,12 @@ begin simpa [hk₁] using sub_one_norm_eq_eval_cyclotomic hζ this hirr, end -/-- If `irreducible (cyclotomic (p ^ (k + 1)) K)` and -`irreducible (cyclotomic (p ^ (k - s + 1)) K))` (in particular for `K = ℚ`) and `p` is a prime, +/-- If `irreducible (cyclotomic (p ^ (k + 1)) K)` (in particular for `K = ℚ`) and `p` is a prime, then the norm of `ζ ^ (p ^ s) - 1` is `p ^ (p ^ s)` if `1 ≤ k`. -/ lemma pow_sub_one_norm_prime_pow_of_one_le [hne : ne_zero ((p : ℕ) : K)] {k s : ℕ} (hζ : is_primitive_root ζ ↑(p ^ (k + 1))) [hpri : fact (p : ℕ).prime] [hcycl : is_cyclotomic_extension {p ^ (k + 1)} K L] - (hirr : irreducible (cyclotomic (↑(p ^ (k + 1)) : ℕ) K)) - (hirr₁ : irreducible (cyclotomic (↑(p ^ (k - s + 1)) : ℕ) K)) (hs : s ≤ k) + (hirr : irreducible (cyclotomic (↑(p ^ (k + 1)) : ℕ) K)) (hs : s ≤ k) (hk : 1 ≤ k) : norm K (ζ ^ ((p : ℕ) ^ s) - 1) = p ^ ((p : ℕ) ^ s) := begin by_cases htwo : p ^ (k - s + 1) = 2, @@ -474,7 +470,7 @@ begin rw [hζ.pow_sub_one_norm_two hirr], rw [hk₁, pow_succ, pow_mul, neg_eq_neg_one_mul, mul_pow, neg_one_sq, one_mul, ← pow_mul, ← pow_succ] }, - { exact hζ.pow_sub_one_norm_prime_pow_ne_two hirr hirr₁ hs htwo } + { exact hζ.pow_sub_one_norm_prime_pow_ne_two hirr hs htwo } end end is_primitive_root @@ -505,14 +501,13 @@ begin exact (zeta_primitive_root n K L).sub_one_norm_is_prime_pow hn hirr h, end -/-- If `irreducible (cyclotomic (p ^ (k + 1)) K)` and `irreducible (cyclotomic (p ^ (k - s + 1)) K)` -(in particular for `K = ℚ`) and `p` is a prime, then the norm of -`(zeta (p ^ (k + 1)) K L) ^ (p ^ s) - 1` is `p ^ (p ^ s)` if `p ^ (k - s + 1) ≠ 2`. -/ +/-- If `irreducible (cyclotomic (p ^ (k + 1)) K)` (in particular for `K = ℚ`) and `p` is a prime, +then the norm of `(zeta (p ^ (k + 1)) K L) ^ (p ^ s) - 1` is `p ^ (p ^ s)` +if `p ^ (k - s + 1) ≠ 2`. -/ lemma prime_ne_two_pow_norm_zeta_pow_sub_one [ne_zero ((p : ℕ) : K)] {k : ℕ} [hpri : fact (p : ℕ).prime] [is_cyclotomic_extension {p ^ (k + 1)} K L] - (hirr : irreducible (cyclotomic (↑(p ^ (k + 1)) : ℕ) K)) {s : ℕ} - (hirr₁ : irreducible (cyclotomic (↑(p ^ (k - s + 1)) : ℕ) K)) (hs : s ≤ k) + (hirr : irreducible (cyclotomic (↑(p ^ (k + 1)) : ℕ) K)) {s : ℕ} (hs : s ≤ k) (htwo : p ^ (k - s + 1) ≠ 2) : norm K ((zeta (p ^ (k + 1)) K L) ^ ((p : ℕ) ^ s) - 1) = p ^ ((p : ℕ) ^ s) := begin @@ -521,7 +516,7 @@ begin { refine ⟨λ hzero, _⟩, rw [pow_coe] at hzero, simpa [ne_zero.ne ((p : ℕ) : L)] using hzero }, - exact (zeta_primitive_root _ K L).pow_sub_one_norm_prime_pow_ne_two hirr hirr₁ hs htwo + exact (zeta_primitive_root _ K L).pow_sub_one_norm_prime_pow_ne_two hirr hs htwo end /-- If `irreducible (cyclotomic (p ^ (k + 1)) K)` (in particular for `K = ℚ`) and `p` is an odd diff --git a/src/ring_theory/polynomial/cyclotomic/basic.lean b/src/ring_theory/polynomial/cyclotomic/basic.lean index c5bf4d7da35c5..1aea80bb1b45f 100644 --- a/src/ring_theory/polynomial/cyclotomic/basic.lean +++ b/src/ring_theory/polynomial/cyclotomic/basic.lean @@ -956,20 +956,28 @@ begin nat.totient_mul_of_prime_of_dvd hp hdiv, mul_comm] } end -/-- If `irreducible (cyclotomic (p ^ n) R)` then `irreducible (cyclotomic p R).` -/ ---TODO: add `irreducible (cyclotomic (p ^ m) R)` if `m ≤ n`. -lemma cyclotomic_irreducible_of_irreducible_pow {p : ℕ} (hp : nat.prime p) - {R} [comm_ring R] [is_domain R] : - ∀ {n : ℕ}, n ≠ 0 → irreducible (cyclotomic (p ^ n) R) → irreducible (cyclotomic p R) -| 0 hn _ := by contradiction -| 1 hn h := by rwa pow_one at h -| (n+2) hn h := +/-- If the `p ^ n`th cyclotomic polynomial is irreducible, so is the `p ^ m`th, for `m ≤ n`. -/ +lemma cyclotomic_irreducible_pow_of_irreducible_pow {p : ℕ} (hp : nat.prime p) + {R} [comm_ring R] [is_domain R] {n m : ℕ} (hmn : m ≤ n) + (h : irreducible (cyclotomic (p ^ n) R)) : irreducible (cyclotomic (p ^ m) R) := begin - rw [pow_succ', ←cyclotomic_expand_eq_cyclotomic hp $ dvd_pow_self p n.succ_ne_zero] at h, - replace h := of_irreducible_expand hp.ne_zero h, - refine cyclotomic_irreducible_of_irreducible_pow n.succ_ne_zero h + unfreezingI + { rcases m.eq_zero_or_pos with rfl | hm, + { simpa using irreducible_X_sub_C (1 : R) }, + obtain ⟨k, rfl⟩ := nat.exists_eq_add_of_le hmn, + induction k with k hk }, + { simpa using h }, + have : m + k ≠ 0 := (add_pos_of_pos_of_nonneg hm k.zero_le).ne', + rw [nat.add_succ, pow_succ', ←cyclotomic_expand_eq_cyclotomic hp $ dvd_pow_self p this] at h, + exact hk (by linarith) (of_irreducible_expand hp.ne_zero h) end +/-- If `irreducible (cyclotomic (p ^ n) R)` then `irreducible (cyclotomic p R).` -/ +lemma cyclotomic_irreducible_of_irreducible_pow {p : ℕ} (hp : nat.prime p) {R} [comm_ring R] + [is_domain R] {n : ℕ} (hn : n ≠ 0) (h : irreducible (cyclotomic (p ^ n) R)) : + irreducible (cyclotomic p R) := +pow_one p ▸ cyclotomic_irreducible_pow_of_irreducible_pow hp hn.bot_lt h + end expand section char_p From 929c90184fb4554bd01224e412d5d8b0b048cd6b Mon Sep 17 00:00:00 2001 From: Haruhisa Enomoto <73296681+haruhisa-enomoto@users.noreply.github.com> Date: Thu, 5 May 2022 13:10:59 +0000 Subject: [PATCH 145/153] refactor(ring_theory/*): Remove unnecessary commutativity assumptions (#13966) This replaces `[comm_ring R]` or `[comm_semiring R]` with `[ring R]` or `[semiring R]`, without changing any proofs. --- src/ring_theory/artinian.lean | 6 +++--- src/ring_theory/finiteness.lean | 2 +- src/ring_theory/ideal/basic.lean | 7 ++++--- src/ring_theory/noetherian.lean | 12 ++++++------ 4 files changed, 14 insertions(+), 13 deletions(-) diff --git a/src/ring_theory/artinian.lean b/src/ring_theory/artinian.lean index 4f32cb64d8db4..dc54afeb8af1b 100644 --- a/src/ring_theory/artinian.lean +++ b/src/ring_theory/artinian.lean @@ -373,7 +373,7 @@ theorem is_artinian_span_of_finite (R) {M} [ring R] [add_comm_group M] [module R [is_artinian_ring R] {A : set M} (hA : finite A) : is_artinian R (submodule.span R A) := is_artinian_of_fg_of_artinian _ (submodule.fg_def.mpr ⟨A, hA, rfl⟩) -theorem is_artinian_ring_of_surjective (R) [comm_ring R] (S) [comm_ring S] +theorem is_artinian_ring_of_surjective (R) [ring R] (S) [ring S] (f : R →+* S) (hf : function.surjective f) [H : is_artinian_ring R] : is_artinian_ring S := begin @@ -381,12 +381,12 @@ begin exact order_embedding.well_founded (ideal.order_embedding_of_surjective f hf) H, end -instance is_artinian_ring_range {R} [comm_ring R] {S} [comm_ring S] (f : R →+* S) +instance is_artinian_ring_range {R} [ring R] {S} [ring S] (f : R →+* S) [is_artinian_ring R] : is_artinian_ring f.range := is_artinian_ring_of_surjective R f.range f.range_restrict f.range_restrict_surjective -theorem is_artinian_ring_of_ring_equiv (R) [comm_ring R] {S} [comm_ring S] +theorem is_artinian_ring_of_ring_equiv (R) [ring R] {S} [ring S] (f : R ≃+* S) [is_artinian_ring R] : is_artinian_ring S := is_artinian_ring_of_surjective R S f.to_ring_hom f.to_equiv.surjective diff --git a/src/ring_theory/finiteness.lean b/src/ring_theory/finiteness.lean index 2f6bd9051d508..9e38911ef024d 100644 --- a/src/ring_theory/finiteness.lean +++ b/src/ring_theory/finiteness.lean @@ -134,7 +134,7 @@ lemma trans {R : Type*} (A B : Type*) [comm_semiring R] [comm_semiring A] [algeb ht, submodule.restrict_scalars_top]⟩⟩ @[priority 100] -- see Note [lower instance priority] -instance finite_type {R : Type*} (A : Type*) [comm_semiring R] [comm_semiring A] +instance finite_type {R : Type*} (A : Type*) [comm_semiring R] [semiring A] [algebra R A] [hRA : finite R A] : algebra.finite_type R A := ⟨subalgebra.fg_of_submodule_fg hRA.1⟩ diff --git a/src/ring_theory/ideal/basic.lean b/src/ring_theory/ideal/basic.lean index 64a34e625a2cf..b085195f21f20 100644 --- a/src/ring_theory/ideal/basic.lean +++ b/src/ring_theory/ideal/basic.lean @@ -15,7 +15,8 @@ import linear_algebra.finsupp # Ideals over a ring -This file defines `ideal R`, the type of ideals over a commutative ring `R`. +This file defines `ideal R`, the type of (left) ideals over a ring `R`. +Note that over commutative rings, left ideals and two-sided ideals are equivalent. ## Implementation notes @@ -23,7 +24,7 @@ This file defines `ideal R`, the type of ideals over a commutative ring `R`. ## TODO -Support one-sided ideals, and ideals over non-commutative rings. +Support right ideals, and two-sided ideals over non-commutative rings. -/ universes u v w @@ -232,7 +233,7 @@ end /-- If P is not properly contained in any maximal ideal then it is not properly contained in any proper ideal -/ -lemma maximal_of_no_maximal {R : Type u} [comm_semiring R] {P : ideal R} +lemma maximal_of_no_maximal {R : Type u} [semiring R] {P : ideal R} (hmax : ∀ m : ideal R, P < m → ¬is_maximal m) (J : ideal R) (hPJ : P < J) : J = ⊤ := begin by_contradiction hnonmax, diff --git a/src/ring_theory/noetherian.lean b/src/ring_theory/noetherian.lean index b66f0037f6514..cb4176840af2d 100644 --- a/src/ring_theory/noetherian.lean +++ b/src/ring_theory/noetherian.lean @@ -288,7 +288,7 @@ begin { rwa [inf_of_le_right (show f.ker ≤ (comap f g.ker), from comap_mono bot_le)] } end -lemma fg_restrict_scalars {R S M : Type*} [comm_ring R] [comm_ring S] [algebra R S] +lemma fg_restrict_scalars {R S M : Type*} [comm_semiring R] [semiring S] [algebra R S] [add_comm_group M] [module S M] [module R M] [is_scalar_tower R S M] (N : submodule S M) (hfin : N.fg) (h : function.surjective (algebra_map R S)) : (submodule.restrict_scalars R N).fg := begin @@ -668,8 +668,8 @@ class is_noetherian_ring (R) [semiring R] extends is_noetherian R R : Prop theorem is_noetherian_ring_iff {R} [semiring R] : is_noetherian_ring R ↔ is_noetherian R R := ⟨λ h, h.1, @is_noetherian_ring.mk _ _⟩ -/-- A commutative ring is Noetherian if and only if all its ideals are finitely-generated. -/ -lemma is_noetherian_ring_iff_ideal_fg (R : Type*) [comm_semiring R] : +/-- A ring is Noetherian if and only if all its ideals are finitely-generated. -/ +lemma is_noetherian_ring_iff_ideal_fg (R : Type*) [semiring R] : is_noetherian_ring R ↔ ∀ I : ideal R, I.fg := is_noetherian_ring_iff.trans is_noetherian_def @@ -756,7 +756,7 @@ theorem is_noetherian_span_of_finite (R) {M} [ring R] [add_comm_group M] [module [is_noetherian_ring R] {A : set M} (hA : finite A) : is_noetherian R (submodule.span R A) := is_noetherian_of_fg_of_noetherian _ (submodule.fg_def.mpr ⟨A, hA, rfl⟩) -theorem is_noetherian_ring_of_surjective (R) [comm_ring R] (S) [comm_ring S] +theorem is_noetherian_ring_of_surjective (R) [ring R] (S) [ring S] (f : R →+* S) (hf : function.surjective f) [H : is_noetherian_ring R] : is_noetherian_ring S := begin @@ -764,12 +764,12 @@ begin exact order_embedding.well_founded (ideal.order_embedding_of_surjective f hf).dual H, end -instance is_noetherian_ring_range {R} [comm_ring R] {S} [comm_ring S] (f : R →+* S) +instance is_noetherian_ring_range {R} [ring R] {S} [ring S] (f : R →+* S) [is_noetherian_ring R] : is_noetherian_ring f.range := is_noetherian_ring_of_surjective R f.range f.range_restrict f.range_restrict_surjective -theorem is_noetherian_ring_of_ring_equiv (R) [comm_ring R] {S} [comm_ring S] +theorem is_noetherian_ring_of_ring_equiv (R) [ring R] {S} [ring S] (f : R ≃+* S) [is_noetherian_ring R] : is_noetherian_ring S := is_noetherian_ring_of_surjective R S f.to_ring_hom f.to_equiv.surjective From 03f5ac93053efd6290745aa21f4dc51addddfd63 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Thu, 5 May 2022 13:11:00 +0000 Subject: [PATCH 146/153] feat(category_theory/simple): simple_iff_subobject_is_simple_order (#13969) Co-authored-by: Scott Morrison --- src/algebra/category/Module/simple.lean | 32 ++++-------- .../limits/shapes/zero_morphisms.lean | 8 +++ src/category_theory/simple.lean | 50 ++++++++++++++++++- .../subobject/factor_thru.lean | 6 +++ src/category_theory/subobject/lattice.lean | 5 ++ 5 files changed, 78 insertions(+), 23 deletions(-) diff --git a/src/algebra/category/Module/simple.lean b/src/algebra/category/Module/simple.lean index fcd07a2862251..7ca97e2ef7a9f 100644 --- a/src/algebra/category/Module/simple.lean +++ b/src/algebra/category/Module/simple.lean @@ -1,41 +1,29 @@ /- Copyright (c) 2022 Pierre-Alexandre Bazin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Pierre-Alexandre Bazin +Authors: Pierre-Alexandre Bazin, Scott Morrison -/ import category_theory.simple import algebra.category.Module.abelian +import algebra.category.Module.subobject import ring_theory.simple_module /-! # Simple objects in the category of `R`-modules -We prove simple modules are simple objects in the category of `R`-modules. -TODO : prove that reciprocally, a simple object in the category of `R`-modules is indeed -a simple module. +We prove simple modules are exactly simple objects in the category of `R`-modules. -/ -section category variables {R M : Type*} [ring R] [add_comm_group M] [module R M] -open category_theory -open Module +open category_theory Module -instance is_simple_module_of [H : is_simple_module R M] : is_simple_module R (of R M) := H +lemma simple_iff_is_simple_module : simple (of R M) ↔ is_simple_module R M := +(simple_iff_subobject_is_simple_order _).trans (subobject_Module (of R M)).is_simple_order_iff /-- A simple module is a simple object in the category of modules. -/ instance simple_of_is_simple_module [is_simple_module R M] : simple (of R M) := -{ mono_is_iso_iff_nonzero := λ N f inj, begin - split, - { unfreezingI { rintro h rfl }, - haveI : unique M := unique_of_epi_zero N, - haveI : nontrivial M := is_simple_module.nontrivial R M, - exact false_of_nontrivial_of_subsingleton M }, - { intro h, - haveI : epi f, - { rw epi_iff_range_eq_top, - refine (eq_bot_or_eq_top f.range).resolve_left _, - exact (mt linear_map.range_eq_bot.mp h)}, - exact is_iso_of_mono_of_epi _ } - end } +simple_iff_is_simple_module.mpr ‹_› -end category +/-- A simple object in the category of modules is a simple module. -/ +instance is_simple_module_of_simple (M : Module R) [simple M] : is_simple_module R M := +simple_iff_is_simple_module.mp (simple.of_iso (of_self_iso M)) diff --git a/src/category_theory/limits/shapes/zero_morphisms.lean b/src/category_theory/limits/shapes/zero_morphisms.lean index 11e19b2693bab..d43911a8b2dae 100644 --- a/src/category_theory/limits/shapes/zero_morphisms.lean +++ b/src/category_theory/limits/shapes/zero_morphisms.lean @@ -351,6 +351,14 @@ def iso_zero_of_epi_zero {X Y : C} (h : epi (0 : X ⟶ Y)) : Y ≅ 0 := inv := 0, hom_inv_id' := (cancel_epi (0 : X ⟶ Y)).mp (by simp) } +/-- If a monomorphism out of `X` is zero, then `X ≅ 0`. -/ +def iso_zero_of_mono_eq_zero {X Y : C} {f : X ⟶ Y} [mono f] (h : f = 0) : X ≅ 0 := +by { unfreezingI { subst h, }, apply iso_zero_of_mono_zero ‹_›, } + +/-- If an epimorphism in to `Y` is zero, then `Y ≅ 0`. -/ +def iso_zero_of_epi_eq_zero {X Y : C} {f : X ⟶ Y} [epi f] (h : f = 0) : Y ≅ 0 := +by { unfreezingI { subst h, }, apply iso_zero_of_epi_zero ‹_›, } + /-- If an object `X` is isomorphic to 0, there's no need to use choice to construct an explicit isomorphism: the zero morphism suffices. -/ def iso_of_is_isomorphic_zero {X : C} (P : is_isomorphic X 0) : X ≅ 0 := diff --git a/src/category_theory/simple.lean b/src/category_theory/simple.lean index 7d5fa5786e259..bbca168d6d05f 100644 --- a/src/category_theory/simple.lean +++ b/src/category_theory/simple.lean @@ -6,6 +6,8 @@ Authors: Markus Himmel, Scott Morrison import category_theory.limits.shapes.zero_morphisms import category_theory.limits.shapes.kernels import category_theory.abelian.basic +import category_theory.subobject.lattice +import order.atoms /-! # Simple objects @@ -114,6 +116,8 @@ by simpa [limits.is_zero.iff_id_eq_zero] using id_nonzero X variable [has_zero_object C] open_locale zero_object +variables (C) + /-- We don't want the definition of 'simple' to include the zero object, so we check that here. -/ lemma zero_not_simple [simple (0 : C)] : false := (simple.mono_is_iso_iff_nonzero (0 : (0 : C) ⟶ (0 : C))).mp ⟨⟨0, by tidy⟩⟩ rfl @@ -176,7 +180,7 @@ end end abelian -section +section indecomposable variables [preadditive C] [has_binary_biproducts C] -- There are another three potential variations of this lemma, @@ -203,6 +207,50 @@ lemma indecomposable_of_simple (X : C) [simple X] : indecomposable X := { apply_instance, }, end⟩ +end indecomposable + +section subobject +variables [has_zero_morphisms C] [has_zero_object C] + +open_locale zero_object +open subobject + +instance {X : C} [simple X] : nontrivial (subobject X) := +⟨⟨mk (0 : 0 ⟶ X), mk (𝟙 X), λ h, begin + haveI := simple.of_iso (iso_of_mk_eq_mk _ _ h), + exact zero_not_simple C, +end⟩⟩ + +instance {X : C} [simple X] : is_simple_order (subobject X) := +{ eq_bot_or_eq_top := begin + rintro ⟨⟨⟨(Y : C), ⟨⟩, (f : Y ⟶ X)⟩, (m : mono f)⟩⟩, resetI, + change mk f = ⊥ ∨ mk f = ⊤, + by_cases h : f = 0, + { exact or.inl (mk_eq_bot_iff_zero.mpr h), }, + { refine or.inr ((is_iso_iff_mk_eq_top _).mp ((simple.mono_is_iso_iff_nonzero f).mpr h)), } +end, } + +/-- If `X` has subobject lattice `{⊥, ⊤}`, then `X` is simple. -/ +lemma simple_of_is_simple_order_subobject (X : C) [is_simple_order (subobject X)] : simple X := +begin + split, introsI, split, + { introI i, + rw subobject.is_iso_iff_mk_eq_top at i, + intro w, + rw ←subobject.mk_eq_bot_iff_zero at w, + exact is_simple_order.bot_ne_top (w.symm.trans i), }, + { intro i, + rcases is_simple_order.eq_bot_or_eq_top (subobject.mk f) with h|h, + { rw subobject.mk_eq_bot_iff_zero at h, + exact false.elim (i h), }, + { exact (subobject.is_iso_iff_mk_eq_top _).mpr h, }, } end +/-- `X` is simple iff it has subobject lattice `{⊥, ⊤}`. -/ +lemma simple_iff_subobject_is_simple_order (X : C) : simple X ↔ is_simple_order (subobject X) := +⟨by { introI h, apply_instance, }, + by { introI h, exact simple_of_is_simple_order_subobject X, }⟩ + +end subobject + end category_theory diff --git a/src/category_theory/subobject/factor_thru.lean b/src/category_theory/subobject/factor_thru.lean index 8133b5c8a4e1e..adf579d3ea797 100644 --- a/src/category_theory/subobject/factor_thru.lean +++ b/src/category_theory/subobject/factor_thru.lean @@ -67,6 +67,8 @@ end (subobject.mk f).factors g ↔ (mono_over.mk' f).factors g := iff.rfl +lemma mk_factors_self (f : X ⟶ Y) [mono f] : (mk f).factors f := ⟨𝟙 _, by simp⟩ + lemma factors_iff {X Y : C} (P : subobject Y) (f : X ⟶ Y) : P.factors f ↔ (representative.obj P).factors f := quot.induction_on P $ λ a, mono_over.factors_congr _ (representative_iso _).symm @@ -108,6 +110,10 @@ classical.some_spec ((factors_iff _ _).mp h) P.factor_thru P.arrow h = 𝟙 P := by { ext, simp, } +@[simp] lemma factor_thru_mk_self (f : X ⟶ Y) [mono f] : + (mk f).factor_thru f (mk_factors_self f) = (underlying_iso f).inv := +by { ext, simp, } + @[simp] lemma factor_thru_comp_arrow {X Y : C} {P : subobject Y} (f : X ⟶ P) (h) : P.factor_thru (f ≫ P.arrow) h = f := by { ext, simp, } diff --git a/src/category_theory/subobject/lattice.lean b/src/category_theory/subobject/lattice.lean index e7b6844b4b69e..359d0877b0f9e 100644 --- a/src/category_theory/subobject/lattice.lean +++ b/src/category_theory/subobject/lattice.lean @@ -297,6 +297,11 @@ zero_of_source_iso_zero _ bot_coe_iso_zero lemma bot_factors_iff_zero {A B : C} (f : A ⟶ B) : (⊥ : subobject B).factors f ↔ f = 0 := ⟨by { rintro ⟨h, rfl⟩, simp }, by { rintro rfl, exact ⟨0, by simp⟩, }⟩ +lemma mk_eq_bot_iff_zero {f : X ⟶ Y} [mono f] : subobject.mk f = ⊥ ↔ f = 0 := +⟨λ h, by simpa [h, bot_factors_iff_zero] using mk_factors_self f, + λ h, mk_eq_mk_of_comm _ _ ((iso_zero_of_mono_eq_zero h).trans has_zero_object.zero_iso_initial) + (by simp [h])⟩ + end zero_order_bot section functor From 420fabf72e27bd2092ae39ee01211538227a7a30 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 5 May 2022 13:11:01 +0000 Subject: [PATCH 147/153] =?UTF-8?q?chore(analysis/normed=5Fspace/exponenti?= =?UTF-8?q?al):=20replace=20`1/x`=20with=20`x=E2=81=BB=C2=B9`=20(#13971)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Note that `one_div` makes `⁻¹` the simp-normal form. --- src/analysis/normed_space/exponential.lean | 35 +++++++++---------- src/analysis/normed_space/spectrum.lean | 14 ++++---- .../special_functions/exponential.lean | 2 +- 3 files changed, 24 insertions(+), 27 deletions(-) diff --git a/src/analysis/normed_space/exponential.lean b/src/analysis/normed_space/exponential.lean index 75c601672c40b..3c449f32ceb5e 100644 --- a/src/analysis/normed_space/exponential.lean +++ b/src/analysis/normed_space/exponential.lean @@ -71,7 +71,7 @@ variables (𝕂 𝔸 : Type*) [field 𝕂] [ring 𝔸] [algebra 𝕂 𝔸] [topo /-- `exp_series 𝕂 𝔸` is the `formal_multilinear_series` whose `n`-th term is the map `(xᵢ) : 𝔸ⁿ ↦ (1/n! : 𝕂) • ∏ xᵢ`. Its sum is the exponential map `exp 𝕂 𝔸 : 𝔸 → 𝔸`. -/ def exp_series : formal_multilinear_series 𝕂 𝔸 𝔸 := -λ n, (1/n! : 𝕂) • continuous_multilinear_map.mk_pi_algebra_fin 𝕂 n 𝔸 +λ n, (n!⁻¹ : 𝕂) • continuous_multilinear_map.mk_pi_algebra_fin 𝕂 n 𝔸 /-- `exp 𝕂 𝔸 : 𝔸 → 𝔸` is the exponential map determined by the action of `𝕂` on `𝔸`. It is defined as the sum of the `formal_multilinear_series` `exp_series 𝕂 𝔸`. -/ @@ -79,17 +79,17 @@ noncomputable def exp (x : 𝔸) : 𝔸 := (exp_series 𝕂 𝔸).sum x variables {𝕂 𝔸} -lemma exp_series_apply_eq (x : 𝔸) (n : ℕ) : exp_series 𝕂 𝔸 n (λ _, x) = (1 / n! : 𝕂) • x^n := +lemma exp_series_apply_eq (x : 𝔸) (n : ℕ) : exp_series 𝕂 𝔸 n (λ _, x) = (n!⁻¹ : 𝕂) • x^n := by simp [exp_series] lemma exp_series_apply_eq' (x : 𝔸) : - (λ n, exp_series 𝕂 𝔸 n (λ _, x)) = (λ n, (1 / n! : 𝕂) • x^n) := + (λ 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, inv_eq_one_div], + rw [div_eq_inv_mul, ←smul_eq_mul], exact exp_series_apply_eq x n, end @@ -97,14 +97,14 @@ lemma exp_series_apply_eq_field' [topological_space 𝕂] [topological_ring 𝕂 (λ 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 : ℕ), (1 / n! : 𝕂) • x^n := +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 : ℕ), (1 / n! : 𝕂) • 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 𝕂] : @@ -113,7 +113,7 @@ funext exp_series_sum_eq_field @[simp] lemma exp_zero [t2_space 𝔸] : exp 𝕂 𝔸 0 = 1 := begin - suffices : (λ x : 𝔸, ∑' (n : ℕ), (1 / n! : 𝕂) • x^n) 0 = ∑' (n : ℕ), if n = 0 then 1 else 0, + suffices : (λ x : 𝔸, ∑' (n : ℕ), (n!⁻¹ : 𝕂) • x^n) 0 = ∑' (n : ℕ), if n = 0 then 1 else 0, { have key : ∀ n ∉ ({0} : finset ℕ), (if n = 0 then (1 : 𝔸) else 0) = 0, from λ n hn, if_neg (finset.not_mem_singleton.mp hn), rw [exp_eq_tsum, this, tsum_eq_sum key, finset.sum_singleton], @@ -153,7 +153,7 @@ lemma norm_exp_series_summable_of_mem_ball (x : 𝔸) lemma norm_exp_series_summable_of_mem_ball' (x : 𝔸) (hx : x ∈ emetric.ball (0 : 𝔸) (exp_series 𝕂 𝔸).radius) : - summable (λ n, ∥(1 / n! : 𝕂) • x^n∥) := + summable (λ n, ∥(n!⁻¹ : 𝕂) • x^n∥) := begin change summable (norm ∘ _), rw ← exp_series_apply_eq', @@ -180,7 +180,7 @@ summable_of_summable_norm (norm_exp_series_summable_of_mem_ball x hx) lemma exp_series_summable_of_mem_ball' (x : 𝔸) (hx : x ∈ emetric.ball (0 : 𝔸) (exp_series 𝕂 𝔸).radius) : - summable (λ n, (1 / n! : 𝕂) • x^n) := + 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 : 𝕂) @@ -194,7 +194,7 @@ formal_multilinear_series.has_sum (exp_series 𝕂 𝔸) hx lemma exp_series_has_sum_exp_of_mem_ball' (x : 𝔸) (hx : x ∈ emetric.ball (0 : 𝔸) (exp_series 𝕂 𝔸).radius) : - has_sum (λ n, (1 / n! : 𝕂) • x^n) (exp 𝕂 𝔸 x):= + has_sum (λ n, (n!⁻¹ : 𝕂) • x^n) (exp 𝕂 𝔸 x):= begin rw ← exp_series_apply_eq', exact exp_series_has_sum_exp_of_mem_ball x hx @@ -339,8 +339,8 @@ begin refine (exp_series 𝕂 𝔸).radius_eq_top_of_summable_norm (λ r, _), refine summable_of_norm_bounded_eventually _ (real.summable_pow_div_factorial r) _, filter_upwards [eventually_cofinite_ne 0] with n hn, - rw [norm_mul, norm_norm (exp_series 𝕂 𝔸 n), exp_series, norm_smul, norm_div, norm_one, norm_pow, - nnreal.norm_eq, norm_eq_abs, abs_cast_nat, mul_comm, ←mul_assoc, ←mul_div_assoc, mul_one], + rw [norm_mul, norm_norm (exp_series 𝕂 𝔸 n), exp_series, norm_smul, norm_inv, norm_pow, + nnreal.norm_eq, norm_eq_abs, abs_cast_nat, mul_comm, ←mul_assoc, ←div_eq_mul_inv], have : ∥continuous_multilinear_map.mk_pi_algebra_fin 𝕂 n 𝔸∥ ≤ 1 := norm_mk_pi_algebra_fin_le_of_pos (nat.pos_of_ne_zero hn), exact mul_le_of_le_one_right (div_nonneg (pow_nonneg r.coe_nonneg n) n!.cast_nonneg) this @@ -359,7 +359,7 @@ 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, ∥(1 / n! : 𝕂) • x^n∥) := +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!∥) := @@ -371,7 +371,7 @@ variables [complete_space 𝔸] lemma exp_series_summable (x : 𝔸) : summable (λ n, exp_series 𝕂 𝔸 n (λ _, x)) := summable_of_summable_norm (norm_exp_series_summable x) -lemma exp_series_summable' (x : 𝔸) : summable (λ n, (1 / n! : 𝕂) • x^n) := +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!) := @@ -380,7 +380,7 @@ 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, (1 / n! : 𝕂) • x^n) (exp 𝕂 𝔸 x):= +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):= @@ -582,10 +582,7 @@ variables (𝕂 𝕂' 𝔸 : Type*) [field 𝕂] [field 𝕂'] [ring 𝔸] [alge `exp_series` on `𝔸`. -/ lemma exp_series_eq_exp_series (n : ℕ) (x : 𝔸) : (exp_series 𝕂 𝔸 n (λ _, x)) = (exp_series 𝕂' 𝔸 n (λ _, x)) := -by rw [exp_series, exp_series, - smul_apply, mk_pi_algebra_fin_apply, list.of_fn_const, list.prod_repeat, - smul_apply, mk_pi_algebra_fin_apply, list.of_fn_const, list.prod_repeat, - one_div, one_div, inv_nat_cast_smul_eq 𝕂 𝕂'] +by rw [exp_series_apply_eq, exp_series_apply_eq, inv_nat_cast_smul_eq 𝕂 𝕂'] /-- If a normed ring `𝔸` is a normed algebra over two fields, then they define the same exponential function on `𝔸`. -/ diff --git a/src/analysis/normed_space/spectrum.lean b/src/analysis/normed_space/spectrum.lean index 6ae2ba0bc0099..3e8b74c74dbad 100644 --- a/src/analysis/normed_space/spectrum.lean +++ b/src/analysis/normed_space/spectrum.lean @@ -378,23 +378,23 @@ begin have hexpmul : exp 𝕜 A a = exp 𝕜 A (a - ↑ₐ z) * ↑ₐ (exp 𝕜 𝕜 z), { rw [algebra_map_exp_comm z, ←exp_add_of_commute (algebra.commutes z (a - ↑ₐz)).symm, sub_add_cancel] }, - let b := ∑' n : ℕ, ((1 / (n + 1).factorial) : 𝕜) • (a - ↑ₐz) ^ n, - have hb : summable (λ n : ℕ, ((1 / (n + 1).factorial) : 𝕜) • (a - ↑ₐz) ^ n), + let b := ∑' n : ℕ, ((n + 1).factorial⁻¹ : 𝕜) • (a - ↑ₐz) ^ n, + have hb : summable (λ n : ℕ, ((n + 1).factorial⁻¹ : 𝕜) • (a - ↑ₐz) ^ n), { refine summable_of_norm_bounded_eventually _ (real.summable_pow_div_factorial ∥a - ↑ₐz∥) _, filter_upwards [filter.eventually_cofinite_ne 0] with n hn, - rw [norm_smul, mul_comm, norm_div, norm_one, is_R_or_C.norm_eq_abs, is_R_or_C.abs_cast_nat, - ←div_eq_mul_one_div], + rw [norm_smul, mul_comm, norm_inv, is_R_or_C.norm_eq_abs, is_R_or_C.abs_cast_nat, + ←div_eq_mul_inv], exact div_le_div (pow_nonneg (norm_nonneg _) n) (norm_pow_le' (a - ↑ₐz) (zero_lt_iff.mpr hn)) (by exact_mod_cast nat.factorial_pos n) (by exact_mod_cast nat.factorial_le (lt_add_one n).le) }, - have h₀ : ∑' n : ℕ, ((1 / (n + 1).factorial) : 𝕜) • (a - ↑ₐz) ^ (n + 1) = (a - ↑ₐz) * b, + have h₀ : ∑' n : ℕ, ((n + 1).factorial⁻¹ : 𝕜) • (a - ↑ₐz) ^ (n + 1) = (a - ↑ₐz) * b, { simpa only [mul_smul_comm, pow_succ] using hb.tsum_mul_left (a - ↑ₐz) }, - have h₁ : ∑' n : ℕ, ((1 / (n + 1).factorial) : 𝕜) • (a - ↑ₐz) ^ (n + 1) = b * (a - ↑ₐz), + have h₁ : ∑' n : ℕ, ((n + 1).factorial⁻¹ : 𝕜) • (a - ↑ₐz) ^ (n + 1) = b * (a - ↑ₐz), { simpa only [pow_succ', algebra.smul_mul_assoc] using hb.tsum_mul_right (a - ↑ₐz) }, have h₃ : exp 𝕜 A (a - ↑ₐz) = 1 + (a - ↑ₐz) * b, { rw exp_eq_tsum, convert tsum_eq_zero_add (exp_series_summable' (a - ↑ₐz)), - simp only [nat.factorial_zero, nat.cast_one, _root_.div_one, pow_zero, one_smul], + simp only [nat.factorial_zero, nat.cast_one, inv_one, pow_zero, one_smul], exact h₀.symm }, rw [spectrum.mem_iff, is_unit.sub_iff, ←one_mul (↑ₐ(exp 𝕜 𝕜 z)), hexpmul, ←_root_.sub_mul, commute.is_unit_mul_iff (algebra.commutes (exp 𝕜 𝕜 z) (exp 𝕜 A (a - ↑ₐz) - 1)).symm, diff --git a/src/analysis/special_functions/exponential.lean b/src/analysis/special_functions/exponential.lean index f5db2d58b1732..7fe7bb71cc15c 100644 --- a/src/analysis/special_functions/exponential.lean +++ b/src/analysis/special_functions/exponential.lean @@ -223,7 +223,7 @@ begin ← 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, - smul_eq_mul, one_div, mul_comm, div_eq_mul_inv] + smul_eq_mul, div_eq_inv_mul] end end real From ec44f459272ea63715a528409dee6d75bcb29c91 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 5 May 2022 14:55:41 +0000 Subject: [PATCH 148/153] feat(data/matrix/basic): even more lemmas about `conj_transpose` and `smul` (#13970) It turns out none of the lemmas in the previous #13938 were the ones I needed. --- src/algebra/star/module.lean | 27 +++++++++++++++++++++++++++ src/data/matrix/basic.lean | 25 +++++++++++++++++++++++++ 2 files changed, 52 insertions(+) diff --git a/src/algebra/star/module.lean b/src/algebra/star/module.lean index d0b67bf4419b4..d6dab39e5d40d 100644 --- a/src/algebra/star/module.lean +++ b/src/algebra/star/module.lean @@ -13,6 +13,8 @@ import linear_algebra.prod We define `star_linear_equiv`, which is the star operation bundled as a star-linear map. It is defined on a star algebra `A` over the base ring `R`. +This file also provides some lemmas that need `algebra.module.basic` imported to prove. + ## TODO - Define `star_linear_equiv` for noncommutative `R`. We only the commutative case for now since, @@ -25,10 +27,35 @@ It is defined on a star algebra `A` over the base ring `R`. equivalence. -/ +section smul_lemmas +variables {R M : Type*} + +@[simp] lemma star_int_cast_smul [ring R] [add_comm_group M] [module R M] [star_add_monoid M] + (n : ℤ) (x : M) : star ((n : R) • x) = (n : R) • star x := +map_int_cast_smul (star_add_equiv : M ≃+ M) R R n x + +@[simp] lemma star_nat_cast_smul [semiring R] [add_comm_monoid M] [module R M] [star_add_monoid M] + (n : ℕ) (x : M) : star ((n : R) • x) = (n : R) • star x := +map_nat_cast_smul (star_add_equiv : M ≃+ M) R R n x + +@[simp] lemma star_inv_int_cast_smul [division_ring R] [add_comm_group M] [module R M] + [star_add_monoid M] (n : ℤ) (x : M) : star ((n⁻¹ : R) • x) = (n⁻¹ : R) • star x := +map_inv_int_cast_smul (star_add_equiv : M ≃+ M) R R n x + +@[simp] lemma star_inv_nat_cast_smul [division_ring R] [add_comm_group M] [module R M] + [star_add_monoid M] (n : ℕ) (x : M) : star ((n⁻¹ : R) • x) = (n⁻¹ : R) • star x := +map_inv_nat_cast_smul (star_add_equiv : M ≃+ M) R R n x + +@[simp] lemma star_rat_cast_smul [division_ring R] [add_comm_group M] [module R M] + [star_add_monoid M] (n : ℚ) (x : M) : star ((n : R) • x) = (n : R) • star x := +map_rat_cast_smul (star_add_equiv : M ≃+ M) _ _ _ x + @[simp] lemma star_rat_smul {R : Type*} [add_comm_group R] [star_add_monoid R] [module ℚ R] (x : R) (n : ℚ) : star (n • x) = n • star x := map_rat_smul (star_add_equiv : R ≃+ R) _ _ +end smul_lemmas + /-- If `A` is a module over a commutative `R` with compatible actions, then `star` is a semilinear equivalence. -/ @[simps] diff --git a/src/data/matrix/basic.lean b/src/data/matrix/basic.lean index e5d4a6f4d63bd..6b32afe45ee3c 100644 --- a/src/data/matrix/basic.lean +++ b/src/data/matrix/basic.lean @@ -1340,7 +1340,12 @@ variants which this lemma would not apply to: * `matrix.conj_transpose_smul_non_comm` * `matrix.conj_transpose_nsmul` * `matrix.conj_transpose_zsmul` +* `matrix.conj_transpose_nat_cast_smul` +* `matrix.conj_transpose_int_cast_smul` +* `matrix.conj_transpose_inv_nat_cast_smul` +* `matrix.conj_transpose_inv_int_cast_smul` * `matrix.conj_transpose_rat_smul` +* `matrix.conj_transpose_rat_cast_smul` -/ @[simp] lemma conj_transpose_smul [has_star R] [has_star α] [has_scalar R α] [star_module R α] (c : R) (M : matrix m n α) : @@ -1365,6 +1370,26 @@ matrix.ext $ by simp (c • M)ᴴ = c • Mᴴ := matrix.ext $ by simp +@[simp] lemma conj_transpose_nat_cast_smul [semiring R] [add_comm_monoid α] + [star_add_monoid α] [module R α] (c : ℕ) (M : matrix m n α) : ((c : R) • M)ᴴ = (c : R) • Mᴴ := +matrix.ext $ by simp + +@[simp] lemma conj_transpose_int_cast_smul [ring R] [add_comm_group α] + [star_add_monoid α] [module R α] (c : ℤ) (M : matrix m n α) : ((c : R) • M)ᴴ = (c : R) • Mᴴ := +matrix.ext $ by simp + +@[simp] lemma conj_transpose_inv_nat_cast_smul [division_ring R] [add_comm_group α] + [star_add_monoid α] [module R α] (c : ℕ) (M : matrix m n α) : ((c : R)⁻¹ • M)ᴴ = (c : R)⁻¹ • Mᴴ := +matrix.ext $ by simp + +@[simp] lemma conj_transpose_inv_int_cast_smul [division_ring R] [add_comm_group α] + [star_add_monoid α] [module R α] (c : ℤ) (M : matrix m n α) : ((c : R)⁻¹ • M)ᴴ = (c : R)⁻¹ • Mᴴ := +matrix.ext $ by simp + +@[simp] lemma conj_transpose_rat_cast_smul [division_ring R] [add_comm_group α] [star_add_monoid α] + [module R α] (c : ℚ) (M : matrix m n α) : ((c : R) • M)ᴴ = (c : R) • Mᴴ := +matrix.ext $ by simp + @[simp] lemma conj_transpose_rat_smul [add_comm_group α] [star_add_monoid α] [module ℚ α] (c : ℚ) (M : matrix m n α) : (c • M)ᴴ = c • Mᴴ := matrix.ext $ by simp From 54af9e9b05a1f98a5817de6da6f8ae433b6863b5 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Thu, 5 May 2022 16:19:12 +0000 Subject: [PATCH 149/153] fix(topology/algebra/infinite_sum): `tsum_neg` doesn't need `summable` (#13950) Both sides are 0 in the not-summable case. --- src/topology/algebra/infinite_sum.lean | 8 ++++++-- 1 file changed, 6 insertions(+), 2 deletions(-) diff --git a/src/topology/algebra/infinite_sum.lean b/src/topology/algebra/infinite_sum.lean index 78292bb753a0e..04a1ebc14cfec 100644 --- a/src/topology/algebra/infinite_sum.lean +++ b/src/topology/algebra/infinite_sum.lean @@ -676,8 +676,12 @@ end section tsum variables [t2_space α] -lemma tsum_neg (hf : summable f) : ∑'b, - f b = - ∑'b, f b := -hf.has_sum.neg.tsum_eq +lemma tsum_neg : ∑'b, - f b = - ∑'b, f b := +begin + by_cases hf : summable f, + { exact hf.has_sum.neg.tsum_eq, }, + { simp [tsum_eq_zero_of_not_summable hf, tsum_eq_zero_of_not_summable (mt summable.of_neg hf)] }, +end lemma tsum_sub (hf : summable f) (hg : summable g) : ∑'b, (f b - g b) = ∑'b, f b - ∑'b, g b := (hf.has_sum.sub hg.has_sum).tsum_eq From 73e5dadd23049c8c5074f3374c2115d52e8d3879 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 5 May 2022 17:30:45 +0000 Subject: [PATCH 150/153] =?UTF-8?q?feat(analysis/special=5Ffunctions/exp):?= =?UTF-8?q?=20add=20limits=20of=20`exp=20z`=20as=20`re=20z=20=E2=86=92=20?= =?UTF-8?q?=C2=B1=E2=88=9E`=20(#13974)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- src/analysis/special_functions/exp.lean | 17 +++++++++++++++++ 1 file changed, 17 insertions(+) diff --git a/src/analysis/special_functions/exp.lean b/src/analysis/special_functions/exp.lean index 16646df652146..1cf7c8cf4394e 100644 --- a/src/analysis/special_functions/exp.lean +++ b/src/analysis/special_functions/exp.lean @@ -260,3 +260,20 @@ by simpa [is_o_iff_tendsto (λ x hx, ((exp_pos x).ne' hx).elim)] using tendsto_div_pow_mul_exp_add_at_top 1 0 n zero_ne_one end real + +namespace complex + +/-- `complex.abs (complex.exp z) → ∞` as `complex.re z → ∞`. TODO: use `bornology.cobounded`. -/ +lemma tendsto_exp_comap_re_at_top : tendsto exp (comap re at_top) (comap abs at_top) := +by simpa only [tendsto_comap_iff, (∘), abs_exp] using real.tendsto_exp_at_top.comp tendsto_comap + +/-- `complex.exp z → 0` as `complex.re z → -∞`.-/ +lemma tendsto_exp_comap_re_at_bot : tendsto exp (comap re at_bot) (𝓝 0) := +tendsto_zero_iff_norm_tendsto_zero.2 $ + by simpa only [norm_eq_abs, abs_exp] using real.tendsto_exp_at_bot.comp tendsto_comap + +lemma tendsto_exp_comap_re_at_bot_nhds_within : tendsto exp (comap re at_bot) (𝓝[≠] 0) := +tendsto_inf.2 ⟨tendsto_exp_comap_re_at_bot, + tendsto_principal.2 $ eventually_of_forall $ exp_ne_zero⟩ + +end complex From c12536a7a1e96489e3dbeb7175ade0c97349df11 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 5 May 2022 19:41:55 +0000 Subject: [PATCH 151/153] fix(gitpod): correct command name (#13976) `leanpkg config` doesn't exist, it's `leanpkg configure`. @b-mehta tricked me in https://github.com/leanprover-community/mathlib/pull/13949#issuecomment-1117589670 --- .gitpod.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.gitpod.yml b/.gitpod.yml index c0cac26eda069..7b6772e69ce87 100644 --- a/.gitpod.yml +++ b/.gitpod.yml @@ -6,5 +6,5 @@ vscode: - jroesch.lean tasks: - - init: leanpkg config && leanproject get-cache --fallback=download-all + - init: leanpkg configure && leanproject get-cache --fallback=download-all command: . /home/gitpod/.profile From 91cc3f063ee3b5985beb6c07af17a93cfb7cbe8d Mon Sep 17 00:00:00 2001 From: Anatole Dedecker Date: Thu, 5 May 2022 20:53:20 +0000 Subject: [PATCH 152/153] feat(linear_algebra/basic): ker of a linear map equals ker of the corresponding group hom (#13858) Co-authored-by: Eric Wieser --- src/linear_algebra/basic.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/src/linear_algebra/basic.lean b/src/linear_algebra/basic.lean index 558762474a163..0ced615b7222f 100644 --- a/src/linear_algebra/basic.lean +++ b/src/linear_algebra/basic.lean @@ -920,6 +920,9 @@ def range [ring_hom_surjective τ₁₂] (f : M →ₛₗ[τ₁₂] M₂) : subm theorem range_coe [ring_hom_surjective τ₁₂] (f : M →ₛₗ[τ₁₂] M₂) : (range f : set M₂) = set.range f := rfl +lemma range_to_add_submonoid [ring_hom_surjective τ₁₂] (f : M →ₛₗ[τ₁₂] M₂) : + f.range.to_add_submonoid = f.to_add_monoid_hom.mrange := rfl + @[simp] theorem mem_range [ring_hom_surjective τ₁₂] {f : M →ₛₗ[τ₁₂] M₂} {x} : x ∈ range f ↔ ∃ y, f y = x := iff.rfl @@ -1004,6 +1007,9 @@ def ker (f : M →ₛₗ[τ₁₂] M₂) : submodule R M := comap f ⊥ @[simp] theorem map_coe_ker (f : M →ₛₗ[τ₁₂] M₂) (x : ker f) : f x = 0 := mem_ker.1 x.2 +lemma ker_to_add_submonoid (f : M →ₛₗ[τ₁₂] M₂) : + f.ker.to_add_submonoid = f.to_add_monoid_hom.mker := rfl + lemma comp_ker_subtype (f : M →ₛₗ[τ₁₂] M₂) : f.comp f.ker.subtype = 0 := linear_map.ext $ λ x, suffices f x = 0, by simp [this], mem_ker.1 x.2 @@ -1116,6 +1122,12 @@ variables {f : M →ₛₗ[τ₁₂] M₂} include R open submodule +lemma range_to_add_subgroup [ring_hom_surjective τ₁₂] (f : M →ₛₗ[τ₁₂] M₂) : + f.range.to_add_subgroup = f.to_add_monoid_hom.range := rfl + +lemma ker_to_add_subgroup (f : M →ₛₗ[τ₁₂] M₂) : + f.ker.to_add_subgroup = f.to_add_monoid_hom.ker := rfl + theorem sub_mem_ker_iff {x y} : x - y ∈ f.ker ↔ f x = f y := by rw [mem_ker, map_sub, sub_eq_zero] From c62dfe61977dd1fc41841cabfcd4d7515c0c6c5d Mon Sep 17 00:00:00 2001 From: Aaron Anderson Date: Thu, 5 May 2022 22:29:30 +0000 Subject: [PATCH 153/153] =?UTF-8?q?feat(model=5Ftheory/skolem):=20Downward?= =?UTF-8?q?=20L=C3=B6wenheim=E2=80=93Skolem=20(#13723)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Proves the Downward Löwenheim–Skolem theorem: If `s` is a set in an `L`-structure `M` and `κ` an infinite cardinal such that `max (# s, L.card) ≤ κ` and `κ ≤ # M`, then `M` has an elementary substructure containing `s` of cardinality `κ`. --- src/model_theory/basic.lean | 23 +++++++-- src/model_theory/skolem.lean | 95 ++++++++++++++++++++++++++++++------ src/model_theory/syntax.lean | 3 ++ 3 files changed, 102 insertions(+), 19 deletions(-) diff --git a/src/model_theory/basic.lean b/src/model_theory/basic.lean index 5ea9cac2e6323..09401033c2b53 100644 --- a/src/model_theory/basic.lean +++ b/src/model_theory/basic.lean @@ -117,9 +117,9 @@ countable_functions.card_functions_le_omega' variables {L} {L' : language.{u' v'}} lemma card_eq_card_functions_add_card_relations : - L.card = cardinal.lift.{v} (# (Σl, L.functions l)) + - cardinal.lift.{u} (# (Σl, L.relations l)) := -by rw [card, symbols, mk_sum] + L.card = cardinal.sum (λ l, (cardinal.lift.{v} (#(L.functions l)))) + + cardinal.sum (λ l, cardinal.lift.{u} (#(L.relations l))) := +by simp [card, symbols] instance [L.is_relational] {n : ℕ} : is_empty (L.functions n) := is_relational.empty_functions n @@ -191,6 +191,23 @@ lemma encodable.countable_functions [h : encodable (Σl, L.functions l)] : L.countable_functions := encodable.countable_functions +@[simp] lemma card_functions_sum (i : ℕ) : + #((L.sum L').functions i) = (#(L.functions i)).lift + cardinal.lift.{u} (#(L'.functions i)) := +by simp [language.sum] + +@[simp] lemma card_relations_sum (i : ℕ) : + #((L.sum L').relations i) = (#(L.relations i)).lift + cardinal.lift.{v} (#(L'.relations i)) := +by simp [language.sum] + +@[simp] lemma card_sum : + (L.sum L').card = cardinal.lift.{max u' v'} L.card + cardinal.lift.{max u v} L'.card := +begin + simp only [card_eq_card_functions_add_card_relations, card_functions_sum, card_relations_sum, + sum_add_distrib', lift_add, lift_sum, lift_lift], + rw [add_assoc, ←add_assoc (cardinal.sum (λ i, (# (L'.functions i)).lift)), + add_comm (cardinal.sum (λ i, (# (L'.functions i)).lift)), add_assoc, add_assoc] +end + variables (L) (M : Type w) /-- A first-order structure on a type `M` consists of interpretations of all the symbols in a given diff --git a/src/model_theory/skolem.lean b/src/model_theory/skolem.lean index f26e7c3eb8180..dcc2bb75a2a6e 100644 --- a/src/model_theory/skolem.lean +++ b/src/model_theory/skolem.lean @@ -12,40 +12,64 @@ import model_theory.elementary_maps * `first_order.language.skolem₁` is a language consisting of Skolem functions for another language. ## Main Results -* `first_order.language.exists_small_elementary_substructure` is a weak version of -Downward Löwenheim–Skolem, showing that any `L`-structure admits a small `L`-elementary -substructure. +* `first_order.language.exists_elementary_substructure_card_eq` is the Downward Löwenheim–Skolem + theorem: If `s` is a set in an `L`-structure `M` and `κ` an infinite cardinal such that + `max (# s, L.card) ≤ κ` and `κ ≤ # M`, then `M` has an elementary substructure containing `s` of + cardinality `κ`. ## TODO -* Bound the cardinality of `L.bounded_formula empty (n + 1)`, and based on that, bound the -cardinality of `(⊥ : (L.sum L.skolem₁).substructure M)` well enough to prove -Downward Löwenheim–Skolem. * Use `skolem₁` recursively to construct an actual Skolemization of a language. -/ universes u v w - - namespace first_order namespace language -open Structure +open Structure cardinal +open_locale cardinal variables (L : language.{u v}) {M : Type w} [nonempty M] [L.Structure M] /-- A language consisting of Skolem functions for another language. Called `skolem₁` because it is the first step in building a Skolemization of a language. -/ -def skolem₁ : language := ⟨λ n, L.bounded_formula empty (n + 1), λ _, empty⟩ +@[simps] def skolem₁ : language := ⟨λ n, L.bounded_formula empty (n + 1), λ _, empty⟩ variables {L} +theorem card_functions_sum_skolem₁ : + # (Σ n, (L.sum L.skolem₁).functions n) = # (Σ n, L.bounded_formula empty (n + 1)) := +begin + simp only [card_functions_sum, skolem₁_functions, lift_id', mk_sigma, sum_add_distrib'], + rw [add_comm, add_eq_max, max_eq_left], + { refine sum_le_sum _ _ (λ n, _), + rw [← lift_le, lift_lift, lift_mk_le], + refine ⟨⟨λ f, (func f default).bd_equal (func f default), λ f g h, _⟩⟩, + rcases h with ⟨rfl, ⟨rfl⟩⟩, + refl }, + { rw ← mk_sigma, + exact infinite_iff.1 (infinite.of_injective (λ n, ⟨n, ⊥⟩) (λ x y xy, (sigma.mk.inj xy).1)) } +end + +theorem card_functions_sum_skolem₁_le : + # (Σ n, (L.sum L.skolem₁).functions n) ≤ max ω L.card := +begin + rw card_functions_sum_skolem₁, + transitivity # (Σ n, L.bounded_formula empty n), + { exact ⟨⟨sigma.map nat.succ (λ _, id), nat.succ_injective.sigma_map + (λ _, function.injective_id)⟩⟩ }, + { refine trans bounded_formula.card_le (lift_le.1 _), + simp only [mk_empty, lift_zero, lift_uzero, zero_add] } +end + /-- The structure assigning each function symbol of `L.skolem₁` to a skolem function generated with choice. -/ noncomputable instance skolem₁_Structure : L.skolem₁.Structure M := ⟨λ n φ x, classical.epsilon (λ a, φ.realize default (fin.snoc x a : _ → M)), λ _ r, empty.elim r⟩ -lemma substructure.skolem₁_reduct_is_elementary (S : (L.sum L.skolem₁).substructure M) : +namespace substructure + +lemma skolem₁_reduct_is_elementary (S : (L.sum L.skolem₁).substructure M) : (Lhom.sum_inl.substructure_reduct S).is_elementary := begin apply (Lhom.sum_inl.substructure_reduct S).is_elementary_of_exists, @@ -56,23 +80,24 @@ begin end /-- Any `L.sum L.skolem₁`-substructure is an elementary `L`-substructure. -/ -noncomputable def substructure.elementary_skolem₁_reduct (S : (L.sum L.skolem₁).substructure M) : +noncomputable def elementary_skolem₁_reduct (S : (L.sum L.skolem₁).substructure M) : L.elementary_substructure M := ⟨Lhom.sum_inl.substructure_reduct S, λ _, S.skolem₁_reduct_is_elementary⟩ -lemma substructure.coe_sort_elementary_skolem₁_reduct +lemma coe_sort_elementary_skolem₁_reduct (S : (L.sum L.skolem₁).substructure M) : (S.elementary_skolem₁_reduct : Type w) = S := rfl -open cardinal -open_locale cardinal +end substructure + +open substructure variables (L) (M) instance : small (⊥ : (L.sum L.skolem₁).substructure M).elementary_skolem₁_reduct := begin - rw [substructure.coe_sort_elementary_skolem₁_reduct], + rw [coe_sort_elementary_skolem₁_reduct], apply_instance, end @@ -80,5 +105,43 @@ theorem exists_small_elementary_substructure : ∃ (S : L.elementary_substructure M), small.{max u v} S := ⟨substructure.elementary_skolem₁_reduct ⊥, infer_instance⟩ +variables {L M} + +/-- The Downward Löwenheim–Skolem Theorem : + If `s` is a set in an `L`-structure `M` and `κ` an infinite cardinal such that + `max (# s, L.card) ≤ κ` and `κ ≤ # M`, then `M` has an elementary substructure containing `s` of + cardinality `κ`. -/ +theorem exists_elementary_substructure_card_eq (s : set M) (κ : cardinal.{max u v w}) + (h1 : ω ≤ κ) + (h2 : cardinal.lift.{max u v} (# s) ≤ κ) + (h3 : cardinal.lift.{w} L.card ≤ κ) + (h4 : κ ≤ cardinal.lift.{max u v} (# M)) : + ∃ (S : L.elementary_substructure M), s ⊆ S ∧ cardinal.lift.{max u v} (# S) = κ := +begin + obtain ⟨s', rfl⟩ := cardinal.le_mk_iff_exists_set.1 h4, + refine ⟨elementary_skolem₁_reduct (closure (L.sum L.skolem₁) + (s ∪ (equiv.ulift.{(max u v) w} '' s'))), + (s.subset_union_left _).trans subset_closure, _⟩, + rw [coe_sort_elementary_skolem₁_reduct], + refine le_antisymm (lift_le.1 _) _, + { rw lift_lift, + refine lift_card_closure_le.trans _, + rw max_le_iff at *, + rw [← lift_le, lift_lift, lift_le, add_comm, add_eq_max, max_le_iff, lift_id'], + { refine ⟨h1, _, _⟩, + { refine (lift_le.2 card_functions_sum_skolem₁_le).trans _, + rw [lift_max', lift_omega, max_le_iff, ← lift_lift, lift_id], + exact ⟨h1, h3⟩, }, + { refine ((lift_le.2 (mk_union_le _ _)).trans _), + rw [lift_add, add_comm, mk_image_eq_lift _ _ equiv.ulift.injective, ← lift_lift, lift_id', + add_eq_max h1, lift_id', max_eq_left h2] } }, + { rw [← lift_lift, lift_id, ← lift_omega, lift_le, ← infinite_iff], + exact infinite.of_injective (λ n, ⟨n, sum.inr bounded_formula.falsum⟩) + (λ x y xy, (sigma.ext_iff.1 xy).1) } }, + { rw [← lift_le, lift_lift, ← mk_image_eq_lift _ s' equiv.ulift.injective, lift_mk_le], + exact ⟨⟨set.inclusion ((set.subset_union_right _ _).trans subset_closure), + set.inclusion_injective _⟩⟩ }, +end + end language end first_order diff --git a/src/model_theory/syntax.lean b/src/model_theory/syntax.lean index 2e7f756661939..41c25e297c725 100644 --- a/src/model_theory/syntax.lean +++ b/src/model_theory/syntax.lean @@ -113,6 +113,9 @@ def functions.apply₂ (f : L.functions 2) (t₁ t₂ : L.term α) : L.term α : namespace term +instance inhabited_of_var [inhabited α] : inhabited (L.term α) := +⟨var default⟩ + instance inhabited_of_constant [inhabited L.constants] : inhabited (L.term α) := ⟨(default : L.constants).term⟩